src/HOL/Nominal/nominal_inductive2.ML
author wenzelm
Thu, 09 Jun 2011 17:51:49 +0200
changeset 44208 47cf4bc789aa
parent 43232 23f352990944
child 44916 2814ff2a6e3e
permissions -rw-r--r--
simplified Name.variant -- discontinued builtin fold_map;
     1 (*  Title:      HOL/Nominal/nominal_inductive2.ML
     2     Author:     Stefan Berghofer, TU Muenchen
     3 
     4 Infrastructure for proving equivariance and strong induction theorems
     5 for inductive predicates involving nominal datatypes.
     6 Experimental version that allows to avoid lists of atoms.
     7 *)
     8 
     9 signature NOMINAL_INDUCTIVE2 =
    10 sig
    11   val prove_strong_ind: string -> string option -> (string * string list) list ->
    12     local_theory -> Proof.state
    13 end
    14 
    15 structure NominalInductive2 : NOMINAL_INDUCTIVE2 =
    16 struct
    17 
    18 val inductive_forall_name = "HOL.induct_forall";
    19 val inductive_forall_def = @{thm induct_forall_def};
    20 val inductive_atomize = @{thms induct_atomize};
    21 val inductive_rulify = @{thms induct_rulify};
    22 
    23 fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify [];
    24 
    25 val atomize_conv =
    26   Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE))
    27     (HOL_basic_ss addsimps inductive_atomize);
    28 val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
    29 fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
    30   (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
    31 
    32 val fresh_postprocess =
    33   Simplifier.full_simplify (HOL_basic_ss addsimps
    34     [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim},
    35      @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]);
    36 
    37 fun preds_of ps t = inter (op = o apfst dest_Free) (Term.add_frees t []) ps;
    38 
    39 val perm_bool = mk_meta_eq @{thm perm_bool};
    40 val perm_boolI = @{thm perm_boolI};
    41 val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
    42   (Drule.strip_imp_concl (cprop_of perm_boolI))));
    43 
    44 fun mk_perm_bool pi th = th RS Drule.cterm_instantiate
    45   [(perm_boolI_pi, pi)] perm_boolI;
    46 
    47 fun mk_perm_bool_simproc names = Simplifier.simproc_global_i
    48   (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss =>
    49     fn Const ("Nominal.perm", _) $ _ $ t =>
    50          if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t))
    51          then SOME perm_bool else NONE
    52      | _ => NONE);
    53 
    54 fun transp ([] :: _) = []
    55   | transp xs = map hd xs :: transp (map tl xs);
    56 
    57 fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
    58       (Const (s, T), ts) => (case strip_type T of
    59         (Ts, Type (tname, _)) =>
    60           (case NominalDatatype.get_nominal_datatype thy tname of
    61              NONE => fold (add_binders thy i) ts bs
    62            | SOME {descr, index, ...} => (case AList.lookup op =
    63                  (#3 (the (AList.lookup op = descr index))) s of
    64                NONE => fold (add_binders thy i) ts bs
    65              | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') =>
    66                  let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs'
    67                  in (add_binders thy i u
    68                    (fold (fn (u, T) =>
    69                       if exists (fn j => j < i) (loose_bnos u) then I
    70                       else AList.map_default op = (T, [])
    71                         (insert op aconv (incr_boundvars (~i) u)))
    72                           cargs1 bs'), cargs2)
    73                  end) cargs (bs, ts ~~ Ts))))
    74       | _ => fold (add_binders thy i) ts bs)
    75     | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs))
    76   | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
    77   | add_binders thy i _ bs = bs;
    78 
    79 fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of
    80       Const (name, _) =>
    81         if member (op =) names name then SOME (f p q) else NONE
    82     | _ => NONE)
    83   | split_conj _ _ _ _ = NONE;
    84 
    85 fun strip_all [] t = t
    86   | strip_all (_ :: xs) (Const (@{const_name All}, _) $ Abs (s, T, t)) = strip_all xs t;
    87 
    88 (*********************************************************************)
    89 (* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
    90 (* or    ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)            *)
    91 (* to    R ... & id (ALL z. P z (pi_1 o ... o pi_n o t))             *)
    92 (* or    id (ALL z. P z (pi_1 o ... o pi_n o t))                     *)
    93 (*                                                                   *)
    94 (* where "id" protects the subformula from simplification            *)
    95 (*********************************************************************)
    96 
    97 fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ =
    98       (case head_of p of
    99          Const (name, _) =>
   100            if member (op =) names name then SOME (HOLogic.mk_conj (p,
   101              Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
   102                (subst_bounds (pis, strip_all pis q))))
   103            else NONE
   104        | _ => NONE)
   105   | inst_conj_all names ps pis t u =
   106       if member (op aconv) ps (head_of u) then
   107         SOME (Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
   108           (subst_bounds (pis, strip_all pis t)))
   109       else NONE
   110   | inst_conj_all _ _ _ _ _ = NONE;
   111 
   112 fun inst_conj_all_tac k = EVERY
   113   [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]),
   114    REPEAT_DETERM_N k (etac allE 1),
   115    simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1];
   116 
   117 fun map_term f t u = (case f t u of
   118       NONE => map_term' f t u | x => x)
   119 and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of
   120       (NONE, NONE) => NONE
   121     | (SOME t'', NONE) => SOME (t'' $ u)
   122     | (NONE, SOME u'') => SOME (t $ u'')
   123     | (SOME t'', SOME u'') => SOME (t'' $ u''))
   124   | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of
   125       NONE => NONE
   126     | SOME t'' => SOME (Abs (s, T, t'')))
   127   | map_term' _ _ _ = NONE;
   128 
   129 (*********************************************************************)
   130 (*         Prove  F[f t]  from  F[t],  where F is monotone           *)
   131 (*********************************************************************)
   132 
   133 fun map_thm ctxt f tac monos opt th =
   134   let
   135     val prop = prop_of th;
   136     fun prove t =
   137       Goal.prove ctxt [] [] t (fn _ =>
   138         EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
   139           REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
   140           REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
   141   in Option.map prove (map_term f prop (the_default prop opt)) end;
   142 
   143 fun abs_params params t =
   144   let val vs =  map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params)
   145   in (list_all (params, t), (rev vs, subst_bounds (vs, t))) end;
   146 
   147 fun inst_params thy (vs, p) th cts =
   148   let val env = Pattern.first_order_match thy (p, prop_of th)
   149     (Vartab.empty, Vartab.empty)
   150   in Thm.instantiate ([],
   151     map (Envir.subst_term env #> cterm_of thy) vs ~~ cts) th
   152   end;
   153 
   154 fun prove_strong_ind s alt_name avoids ctxt =
   155   let
   156     val thy = Proof_Context.theory_of ctxt;
   157     val ({names, ...}, {raw_induct, intrs, elims, ...}) =
   158       Inductive.the_inductive ctxt (Sign.intern_const thy s);
   159     val ind_params = Inductive.params_of raw_induct;
   160     val raw_induct = atomize_induct ctxt raw_induct;
   161     val elims = map (atomize_induct ctxt) elims;
   162     val monos = Inductive.get_monos ctxt;
   163     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
   164     val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of
   165         [] => ()
   166       | xs => error ("Missing equivariance theorem for predicate(s): " ^
   167           commas_quote xs));
   168     val induct_cases = map fst (fst (Rule_Cases.get (the
   169       (Induct.lookup_inductP ctxt (hd names)))));
   170     val induct_cases' = if null induct_cases then replicate (length intrs) ""
   171       else induct_cases;
   172     val ([raw_induct'], ctxt') = Variable.import_terms false [prop_of raw_induct] ctxt;
   173     val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
   174       HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
   175     val ps = map (fst o snd) concls;
   176 
   177     val _ = (case duplicates (op = o pairself fst) avoids of
   178         [] => ()
   179       | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
   180     val _ = (case subtract (op =) induct_cases (map fst avoids) of
   181         [] => ()
   182       | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
   183     fun mk_avoids params name sets =
   184       let
   185         val (_, ctxt') = Proof_Context.add_fixes
   186           (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt;
   187         fun mk s =
   188           let
   189             val t = Syntax.read_term ctxt' s;
   190             val t' = list_abs_free (params, t) |>
   191               funpow (length params) (fn Abs (_, _, t) => t)
   192           in (t', HOLogic.dest_setT (fastype_of t)) end
   193           handle TERM _ =>
   194             error ("Expression " ^ quote s ^ " to be avoided in case " ^
   195               quote name ^ " is not a set type");
   196         fun add_set p [] = [p]
   197           | add_set (t, T) ((u, U) :: ps) =
   198               if T = U then
   199                 let val S = HOLogic.mk_setT T
   200                 in (Const (@{const_name sup}, S --> S --> S) $ u $ t, T) :: ps
   201                 end
   202               else (u, U) :: add_set (t, T) ps
   203       in
   204         fold (mk #> add_set) sets []
   205       end;
   206 
   207     val prems = map (fn (prem, name) =>
   208       let
   209         val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem);
   210         val concl = incr_boundvars 1 (Logic.strip_assums_concl prem);
   211         val params = Logic.strip_params prem
   212       in
   213         (params,
   214          if null avoids then
   215            map (fn (T, ts) => (HOLogic.mk_set T ts, T))
   216              (fold (add_binders thy 0) (prems @ [concl]) [])
   217          else case AList.lookup op = avoids name of
   218            NONE => []
   219          | SOME sets =>
   220              map (apfst (incr_boundvars 1)) (mk_avoids params name sets),
   221          prems, strip_comb (HOLogic.dest_Trueprop concl))
   222       end) (Logic.strip_imp_prems raw_induct' ~~ induct_cases');
   223 
   224     val atomTs = distinct op = (maps (map snd o #2) prems);
   225     val atoms = map (fst o dest_Type) atomTs;
   226     val ind_sort = if null atomTs then HOLogic.typeS
   227       else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn a => Sign.intern_class thy
   228         ("fs_" ^ Long_Name.base_name a)) atoms));
   229     val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt');
   230     val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt';
   231     val fsT = TFree (fs_ctxt_tyname, ind_sort);
   232 
   233     val inductive_forall_def' = Drule.instantiate'
   234       [SOME (ctyp_of thy fsT)] [] inductive_forall_def;
   235 
   236     fun lift_pred' t (Free (s, T)) ts =
   237       list_comb (Free (s, fsT --> T), t :: ts);
   238     val lift_pred = lift_pred' (Bound 0);
   239 
   240     fun lift_prem (t as (f $ u)) =
   241           let val (p, ts) = strip_comb t
   242           in
   243             if member (op =) ps p then
   244               Const (inductive_forall_name,
   245                 (fsT --> HOLogic.boolT) --> HOLogic.boolT) $
   246                   Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
   247             else lift_prem f $ lift_prem u
   248           end
   249       | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t)
   250       | lift_prem t = t;
   251 
   252     fun mk_fresh (x, T) = HOLogic.mk_Trueprop
   253       (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0);
   254 
   255     val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) =>
   256       let
   257         val params' = params @ [("y", fsT)];
   258         val prem = Logic.list_implies
   259           (map mk_fresh sets @
   260            map (fn prem =>
   261              if null (preds_of ps prem) then prem
   262              else lift_prem prem) prems,
   263            HOLogic.mk_Trueprop (lift_pred p ts));
   264       in abs_params params' prem end) prems);
   265 
   266     val ind_vars =
   267       (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~
   268        map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
   269     val ind_Ts = rev (map snd ind_vars);
   270 
   271     val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   272       (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
   273         HOLogic.list_all (ind_vars, lift_pred p
   274           (map (fold_rev (NominalDatatype.mk_perm ind_Ts)
   275             (map Bound (length atomTs downto 1))) ts)))) concls));
   276 
   277     val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   278       (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
   279         lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls));
   280 
   281     val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) =>
   282       map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies
   283           (map_filter (fn prem =>
   284              if null (preds_of ps prem) then SOME prem
   285              else map_term (split_conj (K o I) names) prem prem) prems, q))))
   286         (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
   287            (NominalDatatype.fresh_star_const U T $ u $ t)) sets)
   288              (ts ~~ binder_types (fastype_of p)) @
   289          map (fn (u, U) => HOLogic.mk_Trueprop (Const (@{const_name finite},
   290            HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |>
   291       split_list) prems |> split_list;
   292 
   293     val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp";
   294     val pt2_atoms = map (fn a => Global_Theory.get_thm thy
   295       ("pt_" ^ Long_Name.base_name a ^ "2")) atoms;
   296     val eqvt_ss = Simplifier.global_context thy HOL_basic_ss
   297       addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms)
   298       addsimprocs [mk_perm_bool_simproc ["Fun.id"],
   299         NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun];
   300     val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij";
   301     val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms;
   302     val at_insts = map (NominalAtoms.at_inst_of thy) atoms;
   303     val dj_thms = maps (fn a =>
   304       map (NominalAtoms.dj_thm_of thy a) (remove (op =) a atoms)) atoms;
   305     val finite_ineq = map2 (fn th => fn th' => th' RS (th RS
   306       @{thm pt_set_finite_ineq})) pt_insts at_insts;
   307     val perm_set_forget =
   308       map (fn th => th RS @{thm dj_perm_set_forget}) dj_thms;
   309     val perm_freshs_freshs = atomTs ~~ map2 (fn th => fn th' => th' RS (th RS
   310       @{thm pt_freshs_freshs})) pt_insts at_insts;
   311 
   312     fun obtain_fresh_name ts sets (T, fin) (freshs, ths1, ths2, ths3, ctxt) =
   313       let
   314         val thy = Proof_Context.theory_of ctxt;
   315         (** protect terms to avoid that fresh_star_prod_set interferes with  **)
   316         (** pairs used in introduction rules of inductive predicate          **)
   317         fun protect t =
   318           let val T = fastype_of t in Const ("Fun.id", T --> T) $ t end;
   319         val p = foldr1 HOLogic.mk_prod (map protect ts);
   320         val atom = fst (dest_Type T);
   321         val {at_inst, ...} = NominalAtoms.the_atom_info thy atom;
   322         val fs_atom = Global_Theory.get_thm thy
   323           ("fs_" ^ Long_Name.base_name atom ^ "1");
   324         val avoid_th = Drule.instantiate'
   325           [SOME (ctyp_of thy (fastype_of p))] [SOME (cterm_of thy p)]
   326           ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding});
   327         val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result
   328           (fn _ => EVERY
   329             [rtac avoid_th 1,
   330              full_simp_tac (HOL_ss addsimps [@{thm fresh_star_prod_set}]) 1,
   331              full_simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1,
   332              rotate_tac 1 1,
   333              REPEAT (etac conjE 1)])
   334           [] ctxt;
   335         val (Ts1, _ :: Ts2) = take_prefix (not o equal T) (map snd sets);
   336         val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2);
   337         val (pis1, pis2) = chop (length Ts1)
   338           (map Bound (length pTs - 1 downto 0));
   339         val _ $ (f $ (_ $ pi $ l) $ r) = prop_of th2
   340         val th2' =
   341           Goal.prove ctxt [] []
   342             (list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop
   343                (f $ fold_rev (NominalDatatype.mk_perm (rev pTs))
   344                   (pis1 @ pi :: pis2) l $ r)))
   345             (fn _ => cut_facts_tac [th2] 1 THEN
   346                full_simp_tac (HOL_basic_ss addsimps perm_set_forget) 1) |>
   347           Simplifier.simplify eqvt_ss
   348       in
   349         (freshs @ [term_of cx],
   350          ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt')
   351       end;
   352 
   353     fun mk_ind_proof ctxt' thss =
   354       Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} =>
   355         let val th = Goal.prove ctxt [] [] concl (fn {context, ...} =>
   356           rtac raw_induct 1 THEN
   357           EVERY (maps (fn (((((_, sets, oprems, _),
   358               vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) =>
   359             [REPEAT (rtac allI 1), simp_tac eqvt_ss 1,
   360              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
   361                let
   362                  val (cparams', (pis, z)) =
   363                    chop (length params - length atomTs - 1) (map #2 params) ||>
   364                    (map term_of #> split_last);
   365                  val params' = map term_of cparams'
   366                  val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
   367                  val pi_sets = map (fn (t, _) =>
   368                    fold_rev (NominalDatatype.mk_perm []) pis t) sets';
   369                  val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
   370                  val gprems1 = map_filter (fn (th, t) =>
   371                    if null (preds_of ps t) then SOME th
   372                    else
   373                      map_thm ctxt' (split_conj (K o I) names)
   374                        (etac conjunct1 1) monos NONE th)
   375                    (gprems ~~ oprems);
   376                  val vc_compat_ths' = map2 (fn th => fn p =>
   377                    let
   378                      val th' = gprems1 MRS inst_params thy p th cparams';
   379                      val (h, ts) =
   380                        strip_comb (HOLogic.dest_Trueprop (concl_of th'))
   381                    in
   382                      Goal.prove ctxt' [] []
   383                        (HOLogic.mk_Trueprop (list_comb (h,
   384                           map (fold_rev (NominalDatatype.mk_perm []) pis) ts)))
   385                        (fn _ => simp_tac (HOL_basic_ss addsimps
   386                           (fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1)
   387                    end) vc_compat_ths vc_compat_vs;
   388                  val (vc_compat_ths1, vc_compat_ths2) =
   389                    chop (length vc_compat_ths - length sets) vc_compat_ths';
   390                  val vc_compat_ths1' = map
   391                    (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv
   392                       (Simplifier.rewrite eqvt_ss)))) vc_compat_ths1;
   393                  val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold
   394                    (obtain_fresh_name ts sets)
   395                    (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], ctxt');
   396                  fun concat_perm pi1 pi2 =
   397                    let val T = fastype_of pi1
   398                    in if T = fastype_of pi2 then
   399                        Const ("List.append", T --> T --> T) $ pi1 $ pi2
   400                      else pi2
   401                    end;
   402                  val pis'' = fold_rev (concat_perm #> map) pis' pis;
   403                  val ihyp' = inst_params thy vs_ihypt ihyp
   404                    (map (fold_rev (NominalDatatype.mk_perm [])
   405                       (pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]);
   406                  fun mk_pi th =
   407                    Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
   408                        addsimprocs [NominalDatatype.perm_simproc])
   409                      (Simplifier.simplify eqvt_ss
   410                        (fold_rev (mk_perm_bool o cterm_of thy)
   411                          (pis' @ pis) th));
   412                  val gprems2 = map (fn (th, t) =>
   413                    if null (preds_of ps t) then mk_pi th
   414                    else
   415                      mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis''))
   416                        (inst_conj_all_tac (length pis'')) monos (SOME t) th)))
   417                    (gprems ~~ oprems);
   418                  val perm_freshs_freshs' = map (fn (th, (_, T)) =>
   419                    th RS the (AList.lookup op = perm_freshs_freshs T))
   420                      (fresh_ths2 ~~ sets);
   421                  val th = Goal.prove ctxt'' [] []
   422                    (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
   423                      map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts))))
   424                    (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1] @
   425                      map (fn th => rtac th 1) fresh_ths3 @
   426                      [REPEAT_DETERM_N (length gprems)
   427                        (simp_tac (HOL_basic_ss
   428                           addsimps [inductive_forall_def']
   429                           addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
   430                         resolve_tac gprems2 1)]));
   431                  val final = Goal.prove ctxt'' [] [] (term_of concl)
   432                    (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
   433                      addsimps vc_compat_ths1' @ fresh_ths1 @
   434                        perm_freshs_freshs') 1);
   435                  val final' = Proof_Context.export ctxt'' ctxt' [final];
   436                in resolve_tac final' 1 end) context 1])
   437                  (prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems'')))
   438         in
   439           cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN
   440           REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN
   441             etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN
   442             asm_full_simp_tac (simpset_of ctxt) 1)
   443         end) |>
   444         fresh_postprocess |>
   445         singleton (Proof_Context.export ctxt' ctxt);
   446 
   447   in
   448     ctxt'' |>
   449     Proof.theorem NONE (fn thss => fn ctxt =>
   450       let
   451         val rec_name = space_implode "_" (map Long_Name.base_name names);
   452         val rec_qualified = Binding.qualify false rec_name;
   453         val ind_case_names = Rule_Cases.case_names induct_cases;
   454         val induct_cases' = Inductive.partition_rules' raw_induct
   455           (intrs ~~ induct_cases); 
   456         val thss' = map (map atomize_intr) thss;
   457         val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
   458         val strong_raw_induct =
   459           mk_ind_proof ctxt thss' |> Inductive.rulify;
   460         val strong_induct =
   461           if length names > 1 then
   462             (strong_raw_induct, [ind_case_names, Rule_Cases.consumes 0])
   463           else (strong_raw_induct RSN (2, rev_mp),
   464             [ind_case_names, Rule_Cases.consumes 1]);
   465         val (induct_name, inducts_name) =
   466           case alt_name of
   467             NONE => (rec_qualified (Binding.name "strong_induct"),
   468                      rec_qualified (Binding.name "strong_inducts"))
   469           | SOME s => (Binding.name s, Binding.name (s ^ "s"));
   470         val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note
   471           ((induct_name,
   472             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]);
   473         val strong_inducts =
   474           Project_Rule.projects ctxt' (1 upto length names) strong_induct'
   475       in
   476         ctxt' |>
   477         Local_Theory.note
   478           ((inducts_name,
   479             [Attrib.internal (K ind_case_names),
   480              Attrib.internal (K (Rule_Cases.consumes 1))]),
   481            strong_inducts) |> snd
   482       end)
   483       (map (map (rulify_term thy #> rpair [])) vc_compat)
   484   end;
   485 
   486 
   487 (* outer syntax *)
   488 
   489 val _ =
   490   Outer_Syntax.local_theory_to_proof "nominal_inductive2"
   491     "prove strong induction theorem for inductive predicate involving nominal datatypes"
   492     Keyword.thy_goal
   493     (Parse.xname -- 
   494      Scan.option (Parse.$$$ "(" |-- Parse.!!! (Parse.name --| Parse.$$$ ")")) --
   495      (Scan.optional (Parse.$$$ "avoids" |-- Parse.enum1 "|" (Parse.name --
   496       (Parse.$$$ ":" |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) =>
   497         prove_strong_ind name rule_name avoids));
   498 
   499 end