src/HOL/Nominal/nominal_inductive.ML
author wenzelm
Sat, 06 Oct 2007 16:50:04 +0200
changeset 24867 e5b55d7be9bb
parent 24861 cc669ca5f382
child 25824 f56dd9745d1b
permissions -rw-r--r--
simplified interfaces for outer syntax;
     1 (*  Title:      HOL/Nominal/nominal_inductive.ML
     2     ID:         $Id$
     3     Author:     Stefan Berghofer, TU Muenchen
     4 
     5 Infrastructure for proving equivariance and strong induction theorems
     6 for inductive predicates involving nominal datatypes.
     7 *)
     8 
     9 signature NOMINAL_INDUCTIVE =
    10 sig
    11   val prove_strong_ind: string -> (string * string list) list -> theory -> Proof.state
    12   val prove_eqvt: string -> string list -> theory -> theory
    13 end
    14 
    15 structure NominalInductive : NOMINAL_INDUCTIVE =
    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 = MetaSimplifier.rewrite_term thy inductive_rulify [];
    24 
    25 val atomize_conv =
    26   MetaSimplifier.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.forall_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
    31 
    32 val finite_Un = thm "finite_Un";
    33 val supp_prod = thm "supp_prod";
    34 val fresh_prod = thm "fresh_prod";
    35 
    36 val perm_bool = mk_meta_eq (thm "perm_bool");
    37 val perm_boolI = thm "perm_boolI";
    38 val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
    39   (Drule.strip_imp_concl (cprop_of perm_boolI))));
    40 
    41 fun mk_perm_bool_simproc names = Simplifier.simproc_i
    42   (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss =>
    43     fn Const ("Nominal.perm", _) $ _ $ t =>
    44          if the_default "" (try (head_of #> dest_Const #> fst) t) mem names
    45          then SOME perm_bool else NONE
    46      | _ => NONE);
    47 
    48 val allE_Nil = read_instantiate_sg (the_context()) [("x", "[]")] allE;
    49 
    50 fun transp ([] :: _) = []
    51   | transp xs = map hd xs :: transp (map tl xs);
    52 
    53 fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
    54       (Const (s, T), ts) => (case strip_type T of
    55         (Ts, Type (tname, _)) =>
    56           (case NominalPackage.get_nominal_datatype thy tname of
    57              NONE => fold (add_binders thy i) ts bs
    58            | SOME {descr, index, ...} => (case AList.lookup op =
    59                  (#3 (the (AList.lookup op = descr index))) s of
    60                NONE => fold (add_binders thy i) ts bs
    61              | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') =>
    62                  let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs'
    63                  in (add_binders thy i u
    64                    (fold (fn (u, T) =>
    65                       if exists (fn j => j < i) (loose_bnos u) then I
    66                       else insert (op aconv o pairself fst)
    67                         (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2)
    68                  end) cargs (bs, ts ~~ Ts))))
    69       | _ => fold (add_binders thy i) ts bs)
    70     | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs))
    71   | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
    72   | add_binders thy i _ bs = bs;
    73 
    74 fun split_conj f names (Const ("op &", _) $ p $ q) _ = (case head_of p of
    75       Const (name, _) =>
    76         if name mem names then SOME (f p q) else NONE
    77     | _ => NONE)
    78   | split_conj _ _ _ _ = NONE;
    79 
    80 fun strip_all [] t = t
    81   | strip_all (_ :: xs) (Const ("All", _) $ Abs (s, T, t)) = strip_all xs t;
    82 
    83 (*********************************************************************)
    84 (* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
    85 (* or    ALL pi_1 ... pi_n. P (pi_1 o ... o pi_n o t)                *)
    86 (* to    R ... & id (ALL z. (pi_1 o ... o pi_n o t))                 *)
    87 (* or    id (ALL z. (pi_1 o ... o pi_n o t))                         *)
    88 (*                                                                   *)
    89 (* where "id" protects the subformula from simplification            *)
    90 (*********************************************************************)
    91 
    92 fun inst_conj_all names ps pis (Const ("op &", _) $ p $ q) _ =
    93       (case head_of p of
    94          Const (name, _) =>
    95            if name mem names then SOME (HOLogic.mk_conj (p,
    96              Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
    97                (subst_bounds (pis, strip_all pis q))))
    98            else NONE
    99        | _ => NONE)
   100   | inst_conj_all names ps pis t u =
   101       if member (op aconv) ps (head_of u) then
   102         SOME (Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
   103           (subst_bounds (pis, strip_all pis t)))
   104       else NONE
   105   | inst_conj_all _ _ _ _ _ = NONE;
   106 
   107 fun inst_conj_all_tac k = EVERY
   108   [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]),
   109    REPEAT_DETERM_N k (etac allE 1),
   110    simp_tac (HOL_basic_ss addsimps [id_apply]) 1];
   111 
   112 fun map_term f t u = (case f t u of
   113       NONE => map_term' f t u | x => x)
   114 and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of
   115       (NONE, NONE) => NONE
   116     | (SOME t'', NONE) => SOME (t'' $ u)
   117     | (NONE, SOME u'') => SOME (t $ u'')
   118     | (SOME t'', SOME u'') => SOME (t'' $ u''))
   119   | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of
   120       NONE => NONE
   121     | SOME t'' => SOME (Abs (s, T, t'')))
   122   | map_term' _ _ _ = NONE;
   123 
   124 (*********************************************************************)
   125 (*         Prove  F[f t]  from  F[t],  where F is monotone           *)
   126 (*********************************************************************)
   127 
   128 fun map_thm ctxt f tac monos opt th =
   129   let
   130     val prop = prop_of th;
   131     fun prove t =
   132       Goal.prove ctxt [] [] t (fn _ =>
   133         EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
   134           REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
   135           REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
   136   in Option.map prove (map_term f prop (the_default prop opt)) end;
   137 
   138 fun prove_strong_ind s avoids thy =
   139   let
   140     val ctxt = ProofContext.init thy;
   141     val ({names, ...}, {raw_induct, ...}) =
   142       InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
   143     val raw_induct = atomize_induct ctxt raw_induct;
   144     val monos = InductivePackage.get_monos ctxt;
   145     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
   146     val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
   147         [] => ()
   148       | xs => error ("Missing equivariance theorem for predicate(s): " ^
   149           commas_quote xs));
   150     val induct_cases = map fst (fst (RuleCases.get (the
   151       (Induct.lookup_inductP ctxt (hd names)))));
   152     val raw_induct' = Logic.unvarify (prop_of raw_induct);
   153     val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
   154       HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
   155     val ps = map (fst o snd) concls;
   156 
   157     val _ = (case duplicates (op = o pairself fst) avoids of
   158         [] => ()
   159       | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
   160     val _ = assert_all (null o duplicates op = o snd) avoids
   161       (fn (a, _) => error ("Duplicate variable names for case " ^ quote a));
   162     val _ = (case map fst avoids \\ induct_cases of
   163         [] => ()
   164       | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
   165     val avoids' = map (fn name =>
   166       (name, the_default [] (AList.lookup op = avoids name))) induct_cases;
   167     fun mk_avoids params (name, ps) =
   168       let val k = length params - 1
   169       in map (fn x => case find_index (equal x o fst) params of
   170           ~1 => error ("No such variable in case " ^ quote name ^
   171             " of inductive definition: " ^ quote x)
   172         | i => (Bound (k - i), snd (nth params i))) ps
   173       end;
   174 
   175     val prems = map (fn (prem, avoid) =>
   176       let
   177         val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem);
   178         val concl = incr_boundvars 1 (Logic.strip_assums_concl prem);
   179         val params = Logic.strip_params prem
   180       in
   181         (params,
   182          fold (add_binders thy 0) (prems @ [concl]) [] @
   183            map (apfst (incr_boundvars 1)) (mk_avoids params avoid),
   184          prems, strip_comb (HOLogic.dest_Trueprop concl))
   185       end) (Logic.strip_imp_prems raw_induct' ~~ avoids');
   186 
   187     val atomTs = distinct op = (maps (map snd o #2) prems);
   188     val ind_sort = if null atomTs then HOLogic.typeS
   189       else Sign.certify_sort thy (map (fn T => Sign.intern_class thy
   190         ("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs);
   191     val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
   192     val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
   193     val fsT = TFree (fs_ctxt_tyname, ind_sort);
   194 
   195     val inductive_forall_def' = Drule.instantiate'
   196       [SOME (ctyp_of thy fsT)] [] inductive_forall_def;
   197 
   198     fun lift_pred' t (Free (s, T)) ts =
   199       list_comb (Free (s, fsT --> T), t :: ts);
   200     val lift_pred = lift_pred' (Bound 0);
   201 
   202     fun lift_prem (t as (f $ u)) =
   203           let val (p, ts) = strip_comb t
   204           in
   205             if p mem ps then
   206               Const (inductive_forall_name,
   207                 (fsT --> HOLogic.boolT) --> HOLogic.boolT) $
   208                   Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
   209             else lift_prem f $ lift_prem u
   210           end
   211       | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t)
   212       | lift_prem t = t;
   213 
   214     fun mk_distinct [] = []
   215       | mk_distinct ((x, T) :: xs) = List.mapPartial (fn (y, U) =>
   216           if T = U then SOME (HOLogic.mk_Trueprop
   217             (HOLogic.mk_not (HOLogic.eq_const T $ x $ y)))
   218           else NONE) xs @ mk_distinct xs;
   219 
   220     fun mk_fresh (x, T) = HOLogic.mk_Trueprop
   221       (Const ("Nominal.fresh", T --> fsT --> HOLogic.boolT) $ x $ Bound 0);
   222 
   223     val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) =>
   224       let
   225         val params' = params @ [("y", fsT)];
   226         val prem = Logic.list_implies
   227           (map mk_fresh bvars @ mk_distinct bvars @
   228            map (fn prem =>
   229              if null (term_frees prem inter ps) then prem
   230              else lift_prem prem) prems,
   231            HOLogic.mk_Trueprop (lift_pred p ts));
   232         val vs = map (Var o apfst (rpair 0)) (rename_wrt_term prem params')
   233       in
   234         (list_all (params', prem), (rev vs, subst_bounds (vs, prem)))
   235       end) prems);
   236 
   237     val ind_vars =
   238       (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~
   239        map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
   240     val ind_Ts = rev (map snd ind_vars);
   241 
   242     val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   243       (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
   244         HOLogic.list_all (ind_vars, lift_pred p
   245           (map (fold_rev (NominalPackage.mk_perm ind_Ts)
   246             (map Bound (length atomTs downto 1))) ts)))) concls));
   247 
   248     val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   249       (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
   250         lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls));
   251 
   252     val vc_compat = map (fn (params, bvars, prems, (p, ts)) =>
   253       map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies
   254           (List.mapPartial (fn prem =>
   255              if null (ps inter term_frees prem) then SOME prem
   256              else map_term (split_conj (K o I) names) prem prem) prems, q))))
   257         (mk_distinct bvars @
   258          maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
   259            (Const ("Nominal.fresh", U --> T --> HOLogic.boolT) $ u $ t)) bvars)
   260              (ts ~~ binder_types (fastype_of p)))) prems;
   261 
   262     val perm_pi_simp = PureThy.get_thms thy (Name "perm_pi_simp");
   263     val pt2_atoms = map (fn aT => PureThy.get_thm thy
   264       (Name ("pt_" ^ Sign.base_name (fst (dest_Type aT)) ^ "2"))) atomTs;
   265     val eqvt_ss = HOL_basic_ss addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms)
   266       addsimprocs [mk_perm_bool_simproc ["Fun.id"]];
   267     val fresh_bij = PureThy.get_thms thy (Name "fresh_bij");
   268     val perm_bij = PureThy.get_thms thy (Name "perm_bij");
   269     val fs_atoms = map (fn aT => PureThy.get_thm thy
   270       (Name ("fs_" ^ Sign.base_name (fst (dest_Type aT)) ^ "1"))) atomTs;
   271     val exists_fresh' = PureThy.get_thms thy (Name "exists_fresh'");
   272     val fresh_atm = PureThy.get_thms thy (Name "fresh_atm");
   273     val calc_atm = PureThy.get_thms thy (Name "calc_atm");
   274     val perm_fresh_fresh = PureThy.get_thms thy (Name "perm_fresh_fresh");
   275 
   276     fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) =
   277       let
   278         (** protect terms to avoid that supp_prod interferes with   **)
   279         (** pairs used in introduction rules of inductive predicate **)
   280         fun protect t =
   281           let val T = fastype_of t in Const ("Fun.id", T --> T) $ t end;
   282         val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1);
   283         val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
   284             (HOLogic.exists_const T $ Abs ("x", T,
   285               Const ("Nominal.fresh", T --> fastype_of p --> HOLogic.boolT) $
   286                 Bound 0 $ p)))
   287           (fn _ => EVERY
   288             [resolve_tac exists_fresh' 1,
   289              simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
   290         val (([cx], ths), ctxt') = Obtain.result
   291           (fn _ => EVERY
   292             [etac exE 1,
   293              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
   294              full_simp_tac (HOL_basic_ss addsimps [id_apply]) 1,
   295              REPEAT (etac conjE 1)])
   296           [ex] ctxt
   297       in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
   298 
   299     fun mk_proof thy thss =
   300       let val ctxt = ProofContext.init thy
   301       in Goal.prove_global thy [] prems' concl' (fn ihyps =>
   302         let val th = Goal.prove ctxt [] [] concl (fn {context, ...} =>
   303           rtac raw_induct 1 THEN
   304           EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) =>
   305             [REPEAT (rtac allI 1), simp_tac eqvt_ss 1,
   306              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
   307                let
   308                  val (params', (pis, z)) =
   309                    chop (length params - length atomTs - 1) (map term_of params) ||>
   310                    split_last;
   311                  val bvars' = map
   312                    (fn (Bound i, T) => (nth params' (length params' - i), T)
   313                      | (t, T) => (t, T)) bvars;
   314                  val pi_bvars = map (fn (t, _) =>
   315                    fold_rev (NominalPackage.mk_perm []) pis t) bvars';
   316                  val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
   317                  val (freshs1, freshs2, ctxt'') = fold
   318                    (obtain_fresh_name (ts @ pi_bvars))
   319                    (map snd bvars') ([], [], ctxt');
   320                  val freshs2' = NominalPackage.mk_not_sym freshs2;
   321                  val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1);
   322                  fun concat_perm pi1 pi2 =
   323                    let val T = fastype_of pi1
   324                    in if T = fastype_of pi2 then
   325                        Const ("List.append", T --> T --> T) $ pi1 $ pi2
   326                      else pi2
   327                    end;
   328                  val pis'' = fold (concat_perm #> map) pis' pis;
   329                  val env = Pattern.first_order_match thy (ihypt, prop_of ihyp)
   330                    (Vartab.empty, Vartab.empty);
   331                  val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy))
   332                    (map (Envir.subst_vars env) vs ~~
   333                     map (fold_rev (NominalPackage.mk_perm [])
   334                       (rev pis' @ pis)) params' @ [z])) ihyp;
   335                  fun mk_pi th =
   336                    Simplifier.simplify (HOL_basic_ss addsimps [id_apply]
   337                        addsimprocs [NominalPackage.perm_simproc])
   338                      (Simplifier.simplify eqvt_ss
   339                        (fold_rev (fn pi => fn th' => th' RS Drule.cterm_instantiate
   340                          [(perm_boolI_pi, cterm_of thy pi)] perm_boolI)
   341                            (rev pis' @ pis) th));
   342                  val (gprems1, gprems2) = split_list
   343                    (map (fn (th, t) =>
   344                       if null (term_frees t inter ps) then (SOME th, mk_pi th)
   345                       else
   346                         (map_thm ctxt (split_conj (K o I) names)
   347                            (etac conjunct1 1) monos NONE th,
   348                          mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis''))
   349                            (inst_conj_all_tac (length pis'')) monos (SOME t) th))))
   350                       (gprems ~~ oprems)) |>> List.mapPartial I;
   351                  val vc_compat_ths' = map (fn th =>
   352                    let
   353                      val th' = gprems1 MRS
   354                        Thm.instantiate (Thm.first_order_match
   355                          (Conjunction.mk_conjunction_balanced (cprems_of th),
   356                           Conjunction.mk_conjunction_balanced (map cprop_of gprems1))) th;
   357                      val (bop, lhs, rhs) = (case concl_of th' of
   358                          _ $ (fresh $ lhs $ rhs) =>
   359                            (fn t => fn u => fresh $ t $ u, lhs, rhs)
   360                        | _ $ (_ $ (_ $ lhs $ rhs)) =>
   361                            (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs));
   362                      val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop
   363                          (bop (fold_rev (NominalPackage.mk_perm []) pis lhs)
   364                             (fold_rev (NominalPackage.mk_perm []) pis rhs)))
   365                        (fn _ => simp_tac (HOL_basic_ss addsimps
   366                           (fresh_bij @ perm_bij)) 1 THEN rtac th' 1)
   367                    in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end)
   368                      vc_compat_ths;
   369                  val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths';
   370                  (** Since calc_atm simplifies (pi :: 'a prm) o (x :: 'b) to x **)
   371                  (** we have to pre-simplify the rewrite rules                 **)
   372                  val calc_atm_ss = HOL_ss addsimps calc_atm @
   373                     map (Simplifier.simplify (HOL_ss addsimps calc_atm))
   374                       (vc_compat_ths'' @ freshs2');
   375                  val th = Goal.prove ctxt'' [] []
   376                    (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
   377                      map (fold (NominalPackage.mk_perm []) pis') (tl ts))))
   378                    (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1,
   379                      REPEAT_DETERM_N (nprems_of ihyp - length gprems)
   380                        (simp_tac calc_atm_ss 1),
   381                      REPEAT_DETERM_N (length gprems)
   382                        (simp_tac (HOL_ss
   383                           addsimps inductive_forall_def' :: gprems2
   384                           addsimprocs [NominalPackage.perm_simproc]) 1)]));
   385                  val final = Goal.prove ctxt'' [] [] (term_of concl)
   386                    (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
   387                      addsimps vc_compat_ths'' @ freshs2' @
   388                        perm_fresh_fresh @ fresh_atm) 1);
   389                  val final' = ProofContext.export ctxt'' ctxt' [final];
   390                in resolve_tac final' 1 end) context 1])
   391                  (prems ~~ thss ~~ ihyps ~~ prems'')))
   392         in
   393           cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN
   394           REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN
   395             etac impE 1 THEN atac 1 THEN REPEAT (etac allE_Nil 1) THEN
   396             asm_full_simp_tac (simpset_of thy) 1)
   397         end)
   398       end;
   399 
   400   in
   401     thy |>
   402     ProofContext.init |>
   403     Proof.theorem_i NONE (fn thss => ProofContext.theory (fn thy =>
   404       let
   405         val ctxt = ProofContext.init thy;
   406         val rec_name = space_implode "_" (map Sign.base_name names);
   407         val ind_case_names = RuleCases.case_names induct_cases;
   408         val strong_raw_induct =
   409           mk_proof thy (map (map atomize_intr) thss) |>
   410           InductivePackage.rulify;
   411         val strong_induct =
   412           if length names > 1 then
   413             (strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
   414           else (strong_raw_induct RSN (2, rev_mp),
   415             [ind_case_names, RuleCases.consumes 1]);
   416         val ([strong_induct'], thy') = thy |>
   417           Sign.add_path rec_name |>
   418           PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)];
   419         val strong_inducts =
   420           ProjectRule.projects ctxt (1 upto length names) strong_induct'
   421       in
   422         thy' |>
   423         PureThy.add_thmss [(("strong_inducts", strong_inducts),
   424           [ind_case_names, RuleCases.consumes 1])] |> snd |>
   425         Sign.parent_path
   426       end))
   427       (map (map (rulify_term thy #> rpair [])) vc_compat)
   428   end;
   429 
   430 fun prove_eqvt s xatoms thy =
   431   let
   432     val ctxt = ProofContext.init thy;
   433     val ({names, ...}, {raw_induct, intrs, elims, ...}) =
   434       InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
   435     val raw_induct = atomize_induct ctxt raw_induct;
   436     val elims = map (atomize_induct ctxt) elims;
   437     val intrs = map atomize_intr intrs;
   438     val monos = InductivePackage.get_monos ctxt;
   439     val intrs' = InductivePackage.unpartition_rules intrs
   440       (map (fn (((s, ths), (_, k)), th) =>
   441            (s, ths ~~ InductivePackage.infer_intro_vars th k ths))
   442          (InductivePackage.partition_rules raw_induct intrs ~~
   443           InductivePackage.arities_of raw_induct ~~ elims));
   444     val atoms' = NominalAtoms.atoms_of thy;
   445     val atoms =
   446       if null xatoms then atoms' else
   447       let val atoms = map (Sign.intern_type thy) xatoms
   448       in
   449         (case duplicates op = atoms of
   450              [] => ()
   451            | xs => error ("Duplicate atoms: " ^ commas xs);
   452          case atoms \\ atoms' of
   453              [] => ()
   454            | xs => error ("No such atoms: " ^ commas xs);
   455          atoms)
   456       end;
   457     val perm_pi_simp = PureThy.get_thms thy (Name "perm_pi_simp");
   458     val eqvt_ss = HOL_basic_ss addsimps
   459       (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs
   460       [mk_perm_bool_simproc names];
   461     val t = Logic.unvarify (concl_of raw_induct);
   462     val pi = Name.variant (add_term_names (t, [])) "pi";
   463     val ps = map (fst o HOLogic.dest_imp)
   464       (HOLogic.dest_conj (HOLogic.dest_Trueprop t));
   465     fun eqvt_tac th pi (intr, vs) st =
   466       let
   467         fun eqvt_err s = error
   468           ("Could not prove equivariance for introduction rule\n" ^
   469            Sign.string_of_term (theory_of_thm intr)
   470              (Logic.unvarify (prop_of intr)) ^ "\n" ^ s);
   471         val res = SUBPROOF (fn {prems, params, ...} =>
   472           let
   473             val prems' = map (fn th => the_default th (map_thm ctxt
   474               (split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems;
   475             val prems'' = map (fn th' =>
   476               Simplifier.simplify eqvt_ss (th' RS th)) prems';
   477             val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
   478                map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params)
   479                intr
   480           in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
   481           end) ctxt 1 st
   482       in
   483         case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of
   484           NONE => eqvt_err ("Rule does not match goal\n" ^
   485             Sign.string_of_term (theory_of_thm st) (hd (prems_of st)))
   486         | SOME (th, _) => Seq.single th
   487       end;
   488     val thss = map (fn atom =>
   489       let
   490         val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, [])));
   491         val perm_boolI' = Drule.cterm_instantiate
   492           [(perm_boolI_pi, cterm_of thy pi')] perm_boolI
   493       in map (fn th => zero_var_indexes (th RS mp))
   494         (DatatypeAux.split_conj_thm (Goal.prove_global thy [] []
   495           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p =>
   496             HOLogic.mk_imp (p, list_comb
   497              (apsnd (map (NominalPackage.mk_perm [] pi')) (strip_comb p)))) ps)))
   498           (fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs =>
   499               full_simp_tac eqvt_ss 1 THEN
   500               eqvt_tac perm_boolI' pi' intr_vs) intrs'))))
   501       end) atoms
   502   in
   503     fold (fn (name, ths) =>
   504       Sign.add_path (Sign.base_name name) #>
   505       PureThy.add_thmss [(("eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #>
   506       Sign.parent_path) (names ~~ transp thss) thy
   507   end;
   508 
   509 
   510 (* outer syntax *)
   511 
   512 local structure P = OuterParse and K = OuterKeyword in
   513 
   514 val _ = OuterSyntax.keywords ["avoids"];
   515 
   516 val _ =
   517   OuterSyntax.command "nominal_inductive"
   518     "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
   519     (P.name -- Scan.optional (P.$$$ "avoids" |-- P.and_list1 (P.name --
   520       (P.$$$ ":" |-- Scan.repeat1 P.name))) [] >> (fn (name, avoids) =>
   521         Toplevel.print o Toplevel.theory_to_proof (prove_strong_ind name avoids)));
   522 
   523 val _ =
   524   OuterSyntax.command "equivariance"
   525     "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
   526     (P.name -- Scan.optional (P.$$$ "[" |-- P.list1 P.name --| P.$$$ "]") [] >>
   527       (fn (name, atoms) => Toplevel.theory (prove_eqvt name atoms)));
   528 
   529 end;
   530 
   531 end