src/HOL/Nominal/nominal_package.ML
changeset 26337 44473c957672
parent 26128 fe2d24c26e0c
child 26343 0dd2eab7b296
equal deleted inserted replaced
26336:a0e2b706ce73 26337:44473c957672
   138         val (bT as Type (b, []), _) = dest_permT U
   138         val (bT as Type (b, []), _) = dest_permT U
   139       in if aT mem permTs_of u andalso aT <> bT then
   139       in if aT mem permTs_of u andalso aT <> bT then
   140           let
   140           let
   141             val a' = Sign.base_name a;
   141             val a' = Sign.base_name a;
   142             val b' = Sign.base_name b;
   142             val b' = Sign.base_name b;
   143             val cp = PureThy.get_thm thy (Name ("cp_" ^ a' ^ "_" ^ b' ^ "_inst"));
   143             val cp = dynamic_thm thy ("cp_" ^ a' ^ "_" ^ b' ^ "_inst");
   144             val dj = PureThy.get_thm thy (Name ("dj_" ^ b' ^ "_" ^ a'));
   144             val dj = dynamic_thm thy ("dj_" ^ b' ^ "_" ^ a');
   145             val dj_cp' = [cp, dj] MRS dj_cp;
   145             val dj_cp' = [cp, dj] MRS dj_cp;
   146             val cert = SOME o cterm_of thy
   146             val cert = SOME o cterm_of thy
   147           in
   147           in
   148             SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
   148             SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
   149               [cert t, cert r, cert s] dj_cp'))
   149               [cert t, cert r, cert s] dj_cp'))
   307 
   307 
   308     val _ = warning ("length descr: " ^ string_of_int (length descr));
   308     val _ = warning ("length descr: " ^ string_of_int (length descr));
   309     val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
   309     val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
   310 
   310 
   311     val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
   311     val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
   312     val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def");
   312     val perm_fun_def = dynamic_thm thy2 "perm_fun_def";
   313 
   313 
   314     val unfolded_perm_eq_thms =
   314     val unfolded_perm_eq_thms =
   315       if length descr = length new_type_names then []
   315       if length descr = length new_type_names then []
   316       else map standard (List.drop (split_conj_thm
   316       else map standard (List.drop (split_conj_thm
   317         (Goal.prove_global thy2 [] []
   317         (Goal.prove_global thy2 [] []
   351     (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
   351     (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
   352 
   352 
   353     val _ = warning "perm_append_thms";
   353     val _ = warning "perm_append_thms";
   354 
   354 
   355     (*FIXME: these should be looked up statically*)
   355     (*FIXME: these should be looked up statically*)
   356     val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst");
   356     val at_pt_inst = dynamic_thm thy2 "at_pt_inst";
   357     val pt2 = PureThy.get_thm thy2 (Name "pt2");
   357     val pt2 = dynamic_thm thy2 "pt2";
   358 
   358 
   359     val perm_append_thms = List.concat (map (fn a =>
   359     val perm_append_thms = List.concat (map (fn a =>
   360       let
   360       let
   361         val permT = mk_permT (Type (a, []));
   361         val permT = mk_permT (Type (a, []));
   362         val pi1 = Free ("pi1", permT);
   362         val pi1 = Free ("pi1", permT);
   363         val pi2 = Free ("pi2", permT);
   363         val pi2 = Free ("pi2", permT);
   364         val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
   364         val pt_inst = dynamic_thm thy2 ("pt_" ^ Sign.base_name a ^ "_inst");
   365         val pt2' = pt_inst RS pt2;
   365         val pt2' = pt_inst RS pt2;
   366         val pt2_ax = PureThy.get_thm thy2
   366         val pt2_ax = dynamic_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a);
   367           (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a));
       
   368       in List.take (map standard (split_conj_thm
   367       in List.take (map standard (split_conj_thm
   369         (Goal.prove_global thy2 [] []
   368         (Goal.prove_global thy2 [] []
   370              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   369              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   371                 (map (fn ((s, T), x) =>
   370                 (map (fn ((s, T), x) =>
   372                     let val perm = Const (s, permT --> T --> T)
   371                     let val perm = Const (s, permT --> T --> T)
   384 
   383 
   385     (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
   384     (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
   386 
   385 
   387     val _ = warning "perm_eq_thms";
   386     val _ = warning "perm_eq_thms";
   388 
   387 
   389     val pt3 = PureThy.get_thm thy2 (Name "pt3");
   388     val pt3 = dynamic_thm thy2 "pt3";
   390     val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev");
   389     val pt3_rev = dynamic_thm thy2 "pt3_rev";
   391 
   390 
   392     val perm_eq_thms = List.concat (map (fn a =>
   391     val perm_eq_thms = List.concat (map (fn a =>
   393       let
   392       let
   394         val permT = mk_permT (Type (a, []));
   393         val permT = mk_permT (Type (a, []));
   395         val pi1 = Free ("pi1", permT);
   394         val pi1 = Free ("pi1", permT);
   396         val pi2 = Free ("pi2", permT);
   395         val pi2 = Free ("pi2", permT);
   397         (*FIXME: not robust - better access these theorems using NominalData?*)
   396         (*FIXME: not robust - better access these theorems using NominalData?*)
   398         val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst"));
   397         val at_inst = dynamic_thm thy2 ("at_" ^ Sign.base_name a ^ "_inst");
   399         val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
   398         val pt_inst = dynamic_thm thy2 ("pt_" ^ Sign.base_name a ^ "_inst");
   400         val pt3' = pt_inst RS pt3;
   399         val pt3' = pt_inst RS pt3;
   401         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
   400         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
   402         val pt3_ax = PureThy.get_thm thy2
   401         val pt3_ax = dynamic_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a);
   403           (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a));
       
   404       in List.take (map standard (split_conj_thm
   402       in List.take (map standard (split_conj_thm
   405         (Goal.prove_global thy2 [] [] (Logic.mk_implies
   403         (Goal.prove_global thy2 [] [] (Logic.mk_implies
   406              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
   404              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
   407                 permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
   405                 permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
   408               HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   406               HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   419          length new_type_names)
   417          length new_type_names)
   420       end) atoms);
   418       end) atoms);
   421 
   419 
   422     (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
   420     (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
   423 
   421 
   424     val cp1 = PureThy.get_thm thy2 (Name "cp1");
   422     val cp1 = dynamic_thm thy2 "cp1";
   425     val dj_cp = PureThy.get_thm thy2 (Name "dj_cp");
   423     val dj_cp = dynamic_thm thy2 "dj_cp";
   426     val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose");
   424     val pt_perm_compose = dynamic_thm thy2 "pt_perm_compose";
   427     val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev");
   425     val pt_perm_compose_rev = dynamic_thm thy2 "pt_perm_compose_rev";
   428     val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget");
   426     val dj_perm_perm_forget = dynamic_thm thy2 "dj_perm_perm_forget";
   429 
   427 
   430     fun composition_instance name1 name2 thy =
   428     fun composition_instance name1 name2 thy =
   431       let
   429       let
   432         val name1' = Sign.base_name name1;
   430         val name1' = Sign.base_name name1;
   433         val name2' = Sign.base_name name2;
   431         val name2' = Sign.base_name name2;
   435         val permT1 = mk_permT (Type (name1, []));
   433         val permT1 = mk_permT (Type (name1, []));
   436         val permT2 = mk_permT (Type (name2, []));
   434         val permT2 = mk_permT (Type (name2, []));
   437         val augment = map_type_tfree
   435         val augment = map_type_tfree
   438           (fn (x, S) => TFree (x, cp_class :: S));
   436           (fn (x, S) => TFree (x, cp_class :: S));
   439         val Ts = map (augment o body_type) perm_types;
   437         val Ts = map (augment o body_type) perm_types;
   440         val cp_inst = PureThy.get_thm thy
   438         val cp_inst = dynamic_thm thy ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst");
   441           (Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst"));
       
   442         val simps = simpset_of thy addsimps (perm_fun_def ::
   439         val simps = simpset_of thy addsimps (perm_fun_def ::
   443           (if name1 <> name2 then
   440           (if name1 <> name2 then
   444              let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1'))
   441              let val dj = dynamic_thm thy ("dj_" ^ name2' ^ "_" ^ name1')
   445              in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
   442              in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
   446            else
   443            else
   447              let
   444              let
   448                val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst"));
   445                val at_inst = dynamic_thm thy ("at_" ^ name1' ^ "_inst");
   449                val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst"))
   446                val pt_inst = dynamic_thm thy ("pt_" ^ name1' ^ "_inst");
   450              in
   447              in
   451                [cp_inst RS cp1 RS sym,
   448                [cp_inst RS cp1 RS sym,
   452                 at_inst RS (pt_inst RS pt_perm_compose) RS sym,
   449                 at_inst RS (pt_inst RS pt_perm_compose) RS sym,
   453                 at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
   450                 at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
   454             end))
   451             end))
   572 
   569 
   573     (**** Prove that representing set is closed under permutation ****)
   570     (**** Prove that representing set is closed under permutation ****)
   574 
   571 
   575     val _ = warning "proving closure under permutation...";
   572     val _ = warning "proving closure under permutation...";
   576 
   573 
   577     val abs_perm = PureThy.get_thms thy4 (Name "abs_perm");
   574     val abs_perm = dynamic_thms thy4 "abs_perm";
   578 
   575 
   579     val perm_indnames' = List.mapPartial
   576     val perm_indnames' = List.mapPartial
   580       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
   577       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
   581       (perm_indnames ~~ descr);
   578       (perm_indnames ~~ descr);
   582 
   579 
   657             (EVERY [Class.intro_classes_tac [],
   654             (EVERY [Class.intro_classes_tac [],
   658               rewrite_goals_tac [perm_def],
   655               rewrite_goals_tac [perm_def],
   659               asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
   656               asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
   660               asm_full_simp_tac (simpset_of thy addsimps
   657               asm_full_simp_tac (simpset_of thy addsimps
   661                 [Rep RS perm_closed RS Abs_inverse]) 1,
   658                 [Rep RS perm_closed RS Abs_inverse]) 1,
   662               asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
   659               asm_full_simp_tac (HOL_basic_ss addsimps [dynamic_thm thy
   663                 (Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy)
   660                 ("pt_" ^ Sign.base_name atom ^ "3")]) 1]) thy)
   664         (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
   661         (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
   665            new_type_names ~~ tyvars ~~ perm_closed_thms);
   662            new_type_names ~~ tyvars ~~ perm_closed_thms);
   666 
   663 
   667 
   664 
   668     (** prove that new types are in class cp_<name1>_<name2> **)
   665     (** prove that new types are in class cp_<name1>_<name2> **)
   671 
   668 
   672     fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
   669     fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
   673       let
   670       let
   674         val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
   671         val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
   675         val class = Sign.intern_class thy name;
   672         val class = Sign.intern_class thy name;
   676         val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1
   673         val cp1' = dynamic_thm thy (name ^ "_inst") RS cp1
   677       in fold (fn ((((((Abs_inverse, Rep),
   674       in fold (fn ((((((Abs_inverse, Rep),
   678         perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
   675         perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
   679           AxClass.prove_arity
   676           AxClass.prove_arity
   680             (Sign.intern_type thy name,
   677             (Sign.intern_type thy name,
   681               replicate (length tvs) (classes @ cp_classes), [class])
   678               replicate (length tvs) (classes @ cp_classes), [class])
   921       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
   918       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
   922 
   919 
   923     (** prove injectivity of constructors **)
   920     (** prove injectivity of constructors **)
   924 
   921 
   925     val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
   922     val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
   926     val alpha = PureThy.get_thms thy8 (Name "alpha");
   923     val alpha = dynamic_thms thy8 "alpha";
   927     val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh");
   924     val abs_fresh = dynamic_thms thy8 "abs_fresh";
   928 
   925 
   929     val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
   926     val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
   930       let val T = nth_dtyp' i
   927       let val T = nth_dtyp' i
   931       in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
   928       in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
   932         if null dts then NONE else SOME
   929         if null dts then NONE else SOME
  1067 
  1064 
  1068     val _ = warning "proving finite support for the new datatype";
  1065     val _ = warning "proving finite support for the new datatype";
  1069 
  1066 
  1070     val indnames = DatatypeProp.make_tnames recTs;
  1067     val indnames = DatatypeProp.make_tnames recTs;
  1071 
  1068 
  1072     val abs_supp = PureThy.get_thms thy8 (Name "abs_supp");
  1069     val abs_supp = dynamic_thms thy8 "abs_supp";
  1073     val supp_atm = PureThy.get_thms thy8 (Name "supp_atm");
  1070     val supp_atm = dynamic_thms thy8 "supp_atm";
  1074 
  1071 
  1075     val finite_supp_thms = map (fn atom =>
  1072     val finite_supp_thms = map (fn atom =>
  1076       let val atomT = Type (atom, [])
  1073       let val atomT = Type (atom, [])
  1077       in map standard (List.take
  1074       in map standard (List.take
  1078         (split_conj_thm (Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop
  1075         (split_conj_thm (Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop
  1081                (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
  1078                (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
  1082                (indnames ~~ recTs))))
  1079                (indnames ~~ recTs))))
  1083            (fn _ => indtac dt_induct indnames 1 THEN
  1080            (fn _ => indtac dt_induct indnames 1 THEN
  1084             ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
  1081             ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
  1085               (abs_supp @ supp_atm @
  1082               (abs_supp @ supp_atm @
  1086                PureThy.get_thms thy8 (Name ("fs_" ^ Sign.base_name atom ^ "1")) @
  1083                dynamic_thms thy8 ("fs_" ^ Sign.base_name atom ^ "1") @
  1087                List.concat supp_thms))))),
  1084                List.concat supp_thms))))),
  1088          length new_type_names))
  1085          length new_type_names))
  1089       end) atoms;
  1086       end) atoms;
  1090 
  1087 
  1091     val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
  1088     val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
  1208           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
  1205           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
  1209             (Free (tname, T))))
  1206             (Free (tname, T))))
  1210         (descr'' ~~ recTs ~~ tnames)));
  1207         (descr'' ~~ recTs ~~ tnames)));
  1211 
  1208 
  1212     val fin_set_supp = map (fn Type (s, _) =>
  1209     val fin_set_supp = map (fn Type (s, _) =>
  1213       PureThy.get_thm thy9 (Name ("at_" ^ Sign.base_name s ^ "_inst")) RS
  1210       dynamic_thm thy9 ("at_" ^ Sign.base_name s ^ "_inst") RS
  1214         at_fin_set_supp) dt_atomTs;
  1211         at_fin_set_supp) dt_atomTs;
  1215     val fin_set_fresh = map (fn Type (s, _) =>
  1212     val fin_set_fresh = map (fn Type (s, _) =>
  1216       PureThy.get_thm thy9 (Name ("at_" ^ Sign.base_name s ^ "_inst")) RS
  1213       dynamic_thm thy9 ("at_" ^ Sign.base_name s ^ "_inst") RS
  1217         at_fin_set_fresh) dt_atomTs;
  1214         at_fin_set_fresh) dt_atomTs;
  1218     val pt1_atoms = map (fn Type (s, _) =>
  1215     val pt1_atoms = map (fn Type (s, _) =>
  1219       PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "1"))) dt_atomTs;
  1216       dynamic_thm thy9 ("pt_" ^ Sign.base_name s ^ "1")) dt_atomTs;
  1220     val pt2_atoms = map (fn Type (s, _) =>
  1217     val pt2_atoms = map (fn Type (s, _) =>
  1221       PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "2")) RS sym) dt_atomTs;
  1218       dynamic_thm thy9 ("pt_" ^ Sign.base_name s ^ "2") RS sym) dt_atomTs;
  1222     val exists_fresh' = PureThy.get_thms thy9 (Name "exists_fresh'");
  1219     val exists_fresh' = dynamic_thms thy9 "exists_fresh'";
  1223     val fs_atoms = PureThy.get_thms thy9 (Name "fin_supp");
  1220     val fs_atoms = dynamic_thms thy9 "fin_supp";
  1224     val abs_supp = PureThy.get_thms thy9 (Name "abs_supp");
  1221     val abs_supp = dynamic_thms thy9 "abs_supp";
  1225     val perm_fresh_fresh = PureThy.get_thms thy9 (Name "perm_fresh_fresh");
  1222     val perm_fresh_fresh = dynamic_thms thy9 "perm_fresh_fresh";
  1226     val calc_atm = PureThy.get_thms thy9 (Name "calc_atm");
  1223     val calc_atm = dynamic_thms thy9 "calc_atm";
  1227     val fresh_atm = PureThy.get_thms thy9 (Name "fresh_atm");
  1224     val fresh_atm = dynamic_thms thy9 "fresh_atm";
  1228     val fresh_left = PureThy.get_thms thy9 (Name "fresh_left");
  1225     val fresh_left = dynamic_thms thy9 "fresh_left";
  1229     val perm_swap = PureThy.get_thms thy9 (Name "perm_swap");
  1226     val perm_swap = dynamic_thms thy9 "perm_swap";
  1230 
  1227 
  1231     fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
  1228     fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
  1232       let
  1229       let
  1233         val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
  1230         val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
  1234         val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
  1231         val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
  1497       PureThy.hide_thms true [NameSpace.append
  1494       PureThy.hide_thms true [NameSpace.append
  1498         (Sign.full_name thy10 big_rec_name) "induct"];
  1495         (Sign.full_name thy10 big_rec_name) "induct"];
  1499 
  1496 
  1500     (** equivariance **)
  1497     (** equivariance **)
  1501 
  1498 
  1502     val fresh_bij = PureThy.get_thms thy11 (Name "fresh_bij");
  1499     val fresh_bij = dynamic_thms thy11 "fresh_bij";
  1503     val perm_bij = PureThy.get_thms thy11 (Name "perm_bij");
  1500     val perm_bij = dynamic_thms thy11 "perm_bij";
  1504 
  1501 
  1505     val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
  1502     val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
  1506       let
  1503       let
  1507         val permT = mk_permT aT;
  1504         val permT = mk_permT aT;
  1508         val pi = Free ("pi", permT);
  1505         val pi = Free ("pi", permT);
  1536     (** finite support **)
  1533     (** finite support **)
  1537 
  1534 
  1538     val rec_fin_supp_thms = map (fn aT =>
  1535     val rec_fin_supp_thms = map (fn aT =>
  1539       let
  1536       let
  1540         val name = Sign.base_name (fst (dest_Type aT));
  1537         val name = Sign.base_name (fst (dest_Type aT));
  1541         val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1"));
  1538         val fs_name = dynamic_thm thy11 ("fs_" ^ name ^ "1");
  1542         val aset = HOLogic.mk_setT aT;
  1539         val aset = HOLogic.mk_setT aT;
  1543         val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
  1540         val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
  1544         val fins = map (fn (f, T) => HOLogic.mk_Trueprop
  1541         val fins = map (fn (f, T) => HOLogic.mk_Trueprop
  1545           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
  1542           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
  1546             (rec_fns ~~ rec_fn_Ts)
  1543             (rec_fns ~~ rec_fn_Ts)
  1569            (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
  1566            (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
  1570 
  1567 
  1571     val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
  1568     val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
  1572       let
  1569       let
  1573         val name = Sign.base_name (fst (dest_Type aT));
  1570         val name = Sign.base_name (fst (dest_Type aT));
  1574         val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1"));
  1571         val fs_name = dynamic_thm thy11 ("fs_" ^ name ^ "1");
  1575         val a = Free ("a", aT);
  1572         val a = Free ("a", aT);
  1576         val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
  1573         val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
  1577           (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
  1574           (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
  1578       in
  1575       in
  1579         map (fn (((T, U), R), eqvt_th) =>
  1576         map (fn (((T, U), R), eqvt_th) =>