src/HOL/Nominal/nominal_inductive.ML
author wenzelm
Wed, 25 Jun 2008 17:38:32 +0200
changeset 27353 71c4dd53d4cb
parent 27352 8adeff7fd4bc
child 27449 4880da911af0
permissions -rw-r--r--
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
berghofe@22313
     1
(*  Title:      HOL/Nominal/nominal_inductive.ML
berghofe@22313
     2
    ID:         $Id$
berghofe@22313
     3
    Author:     Stefan Berghofer, TU Muenchen
berghofe@22313
     4
berghofe@22530
     5
Infrastructure for proving equivariance and strong induction theorems
berghofe@22530
     6
for inductive predicates involving nominal datatypes.
berghofe@22313
     7
*)
berghofe@22313
     8
berghofe@22313
     9
signature NOMINAL_INDUCTIVE =
berghofe@22313
    10
sig
berghofe@22730
    11
  val prove_strong_ind: string -> (string * string list) list -> theory -> Proof.state
berghofe@22730
    12
  val prove_eqvt: string -> string list -> theory -> theory
berghofe@22313
    13
end
berghofe@22313
    14
berghofe@22313
    15
structure NominalInductive : NOMINAL_INDUCTIVE =
berghofe@22313
    16
struct
berghofe@22313
    17
berghofe@24570
    18
val inductive_forall_name = "HOL.induct_forall";
berghofe@24570
    19
val inductive_forall_def = thm "induct_forall_def";
berghofe@24570
    20
val inductive_atomize = thms "induct_atomize";
berghofe@24570
    21
val inductive_rulify = thms "induct_rulify";
berghofe@24570
    22
berghofe@24570
    23
fun rulify_term thy = MetaSimplifier.rewrite_term thy inductive_rulify [];
berghofe@24570
    24
berghofe@24570
    25
val atomize_conv =
berghofe@24570
    26
  MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE))
berghofe@24570
    27
    (HOL_basic_ss addsimps inductive_atomize);
berghofe@24570
    28
val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
wenzelm@24832
    29
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
wenzelm@26568
    30
  (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
berghofe@24570
    31
berghofe@22530
    32
val finite_Un = thm "finite_Un";
berghofe@22530
    33
val supp_prod = thm "supp_prod";
berghofe@22530
    34
val fresh_prod = thm "fresh_prod";
berghofe@22530
    35
berghofe@24570
    36
val perm_bool = mk_meta_eq (thm "perm_bool");
berghofe@22313
    37
val perm_boolI = thm "perm_boolI";
berghofe@22313
    38
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
berghofe@22313
    39
  (Drule.strip_imp_concl (cprop_of perm_boolI))));
berghofe@22313
    40
berghofe@25824
    41
fun mk_perm_bool pi th = th RS Drule.cterm_instantiate
berghofe@25824
    42
  [(perm_boolI_pi, pi)] perm_boolI;
berghofe@25824
    43
berghofe@24570
    44
fun mk_perm_bool_simproc names = Simplifier.simproc_i
berghofe@24570
    45
  (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss =>
berghofe@24570
    46
    fn Const ("Nominal.perm", _) $ _ $ t =>
berghofe@24570
    47
         if the_default "" (try (head_of #> dest_Const #> fst) t) mem names
berghofe@24570
    48
         then SOME perm_bool else NONE
berghofe@24570
    49
     | _ => NONE);
berghofe@24570
    50
berghofe@22313
    51
fun transp ([] :: _) = []
berghofe@22313
    52
  | transp xs = map hd xs :: transp (map tl xs);
berghofe@22313
    53
berghofe@22530
    54
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
berghofe@22530
    55
      (Const (s, T), ts) => (case strip_type T of
berghofe@22530
    56
        (Ts, Type (tname, _)) =>
berghofe@22530
    57
          (case NominalPackage.get_nominal_datatype thy tname of
berghofe@22530
    58
             NONE => fold (add_binders thy i) ts bs
berghofe@22530
    59
           | SOME {descr, index, ...} => (case AList.lookup op =
berghofe@22530
    60
                 (#3 (the (AList.lookup op = descr index))) s of
berghofe@22530
    61
               NONE => fold (add_binders thy i) ts bs
berghofe@22530
    62
             | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') =>
berghofe@22530
    63
                 let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs'
berghofe@22530
    64
                 in (add_binders thy i u
berghofe@22530
    65
                   (fold (fn (u, T) =>
berghofe@22530
    66
                      if exists (fn j => j < i) (loose_bnos u) then I
berghofe@22530
    67
                      else insert (op aconv o pairself fst)
berghofe@22530
    68
                        (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2)
berghofe@22530
    69
                 end) cargs (bs, ts ~~ Ts))))
berghofe@22530
    70
      | _ => fold (add_binders thy i) ts bs)
berghofe@22530
    71
    | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs))
berghofe@22530
    72
  | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
berghofe@22530
    73
  | add_binders thy i _ bs = bs;
berghofe@22530
    74
berghofe@24570
    75
fun split_conj f names (Const ("op &", _) $ p $ q) _ = (case head_of p of
berghofe@24570
    76
      Const (name, _) =>
berghofe@24570
    77
        if name mem names then SOME (f p q) else NONE
berghofe@24570
    78
    | _ => NONE)
berghofe@24570
    79
  | split_conj _ _ _ _ = NONE;
berghofe@24570
    80
berghofe@24570
    81
fun strip_all [] t = t
berghofe@24570
    82
  | strip_all (_ :: xs) (Const ("All", _) $ Abs (s, T, t)) = strip_all xs t;
berghofe@24570
    83
berghofe@24570
    84
(*********************************************************************)
berghofe@24570
    85
(* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
berghofe@24570
    86
(* or    ALL pi_1 ... pi_n. P (pi_1 o ... o pi_n o t)                *)
berghofe@24570
    87
(* to    R ... & id (ALL z. (pi_1 o ... o pi_n o t))                 *)
berghofe@24570
    88
(* or    id (ALL z. (pi_1 o ... o pi_n o t))                         *)
berghofe@24570
    89
(*                                                                   *)
berghofe@24570
    90
(* where "id" protects the subformula from simplification            *)
berghofe@24570
    91
(*********************************************************************)
berghofe@24570
    92
berghofe@24570
    93
fun inst_conj_all names ps pis (Const ("op &", _) $ p $ q) _ =
berghofe@24570
    94
      (case head_of p of
berghofe@24570
    95
         Const (name, _) =>
berghofe@24570
    96
           if name mem names then SOME (HOLogic.mk_conj (p,
berghofe@24570
    97
             Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
berghofe@24570
    98
               (subst_bounds (pis, strip_all pis q))))
berghofe@24570
    99
           else NONE
berghofe@24570
   100
       | _ => NONE)
berghofe@24570
   101
  | inst_conj_all names ps pis t u =
berghofe@24570
   102
      if member (op aconv) ps (head_of u) then
berghofe@24570
   103
        SOME (Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $
berghofe@24570
   104
          (subst_bounds (pis, strip_all pis t)))
berghofe@24570
   105
      else NONE
berghofe@24570
   106
  | inst_conj_all _ _ _ _ _ = NONE;
berghofe@24570
   107
berghofe@24570
   108
fun inst_conj_all_tac k = EVERY
berghofe@24570
   109
  [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]),
berghofe@24570
   110
   REPEAT_DETERM_N k (etac allE 1),
haftmann@26359
   111
   simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1];
berghofe@24570
   112
berghofe@24570
   113
fun map_term f t u = (case f t u of
berghofe@24570
   114
      NONE => map_term' f t u | x => x)
berghofe@24570
   115
and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of
berghofe@24570
   116
      (NONE, NONE) => NONE
berghofe@24570
   117
    | (SOME t'', NONE) => SOME (t'' $ u)
berghofe@24570
   118
    | (NONE, SOME u'') => SOME (t $ u'')
berghofe@24570
   119
    | (SOME t'', SOME u'') => SOME (t'' $ u''))
berghofe@24570
   120
  | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of
berghofe@24570
   121
      NONE => NONE
berghofe@24570
   122
    | SOME t'' => SOME (Abs (s, T, t'')))
berghofe@24570
   123
  | map_term' _ _ _ = NONE;
berghofe@24570
   124
berghofe@24570
   125
(*********************************************************************)
berghofe@24570
   126
(*         Prove  F[f t]  from  F[t],  where F is monotone           *)
berghofe@24570
   127
(*********************************************************************)
berghofe@24570
   128
berghofe@24570
   129
fun map_thm ctxt f tac monos opt th =
berghofe@24570
   130
  let
berghofe@24570
   131
    val prop = prop_of th;
berghofe@24570
   132
    fun prove t =
berghofe@24570
   133
      Goal.prove ctxt [] [] t (fn _ =>
berghofe@24570
   134
        EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
berghofe@24570
   135
          REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
berghofe@24570
   136
          REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
berghofe@24570
   137
  in Option.map prove (map_term f prop (the_default prop opt)) end;
berghofe@24570
   138
berghofe@27352
   139
val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion;
berghofe@27352
   140
berghofe@25824
   141
fun first_order_matchs pats objs = Thm.first_order_match
berghofe@27352
   142
  (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats),
berghofe@27352
   143
   eta_contract_cterm (Conjunction.mk_conjunction_balanced objs));
berghofe@25824
   144
berghofe@25824
   145
fun first_order_mrs ths th = ths MRS
berghofe@25824
   146
  Thm.instantiate (first_order_matchs (cprems_of th) (map cprop_of ths)) th;
berghofe@25824
   147
berghofe@22730
   148
fun prove_strong_ind s avoids thy =
berghofe@22313
   149
  let
berghofe@22313
   150
    val ctxt = ProofContext.init thy;
berghofe@25824
   151
    val ({names, ...}, {raw_induct, intrs, elims, ...}) =
berghofe@22730
   152
      InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
berghofe@25824
   153
    val ind_params = InductivePackage.params_of raw_induct;
wenzelm@24832
   154
    val raw_induct = atomize_induct ctxt raw_induct;
berghofe@25824
   155
    val elims = map (atomize_induct ctxt) elims;
berghofe@24570
   156
    val monos = InductivePackage.get_monos ctxt;
urbanc@24571
   157
    val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
berghofe@22788
   158
    val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
berghofe@22788
   159
        [] => ()
berghofe@22788
   160
      | xs => error ("Missing equivariance theorem for predicate(s): " ^
berghofe@22788
   161
          commas_quote xs));
berghofe@22530
   162
    val induct_cases = map fst (fst (RuleCases.get (the
wenzelm@24861
   163
      (Induct.lookup_inductP ctxt (hd names)))));
berghofe@22530
   164
    val raw_induct' = Logic.unvarify (prop_of raw_induct);
berghofe@25824
   165
    val elims' = map (Logic.unvarify o prop_of) elims;
berghofe@22530
   166
    val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
berghofe@22530
   167
      HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
berghofe@22530
   168
    val ps = map (fst o snd) concls;
berghofe@22530
   169
berghofe@22530
   170
    val _ = (case duplicates (op = o pairself fst) avoids of
berghofe@22530
   171
        [] => ()
berghofe@22530
   172
      | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
berghofe@22530
   173
    val _ = assert_all (null o duplicates op = o snd) avoids
berghofe@22530
   174
      (fn (a, _) => error ("Duplicate variable names for case " ^ quote a));
berghofe@22530
   175
    val _ = (case map fst avoids \\ induct_cases of
berghofe@22530
   176
        [] => ()
berghofe@22530
   177
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
berghofe@25824
   178
    val avoids' = if null induct_cases then replicate (length intrs) ("", [])
berghofe@25824
   179
      else map (fn name =>
berghofe@25824
   180
        (name, the_default [] (AList.lookup op = avoids name))) induct_cases;
berghofe@22530
   181
    fun mk_avoids params (name, ps) =
berghofe@22530
   182
      let val k = length params - 1
berghofe@22530
   183
      in map (fn x => case find_index (equal x o fst) params of
berghofe@22530
   184
          ~1 => error ("No such variable in case " ^ quote name ^
berghofe@22530
   185
            " of inductive definition: " ^ quote x)
berghofe@22530
   186
        | i => (Bound (k - i), snd (nth params i))) ps
berghofe@22530
   187
      end;
berghofe@22530
   188
berghofe@22530
   189
    val prems = map (fn (prem, avoid) =>
berghofe@22530
   190
      let
berghofe@22530
   191
        val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem);
berghofe@22530
   192
        val concl = incr_boundvars 1 (Logic.strip_assums_concl prem);
berghofe@22530
   193
        val params = Logic.strip_params prem
berghofe@22530
   194
      in
berghofe@22530
   195
        (params,
berghofe@22530
   196
         fold (add_binders thy 0) (prems @ [concl]) [] @
berghofe@22530
   197
           map (apfst (incr_boundvars 1)) (mk_avoids params avoid),
berghofe@22530
   198
         prems, strip_comb (HOLogic.dest_Trueprop concl))
berghofe@22530
   199
      end) (Logic.strip_imp_prems raw_induct' ~~ avoids');
berghofe@22530
   200
berghofe@22530
   201
    val atomTs = distinct op = (maps (map snd o #2) prems);
berghofe@22530
   202
    val ind_sort = if null atomTs then HOLogic.typeS
berghofe@22530
   203
      else Sign.certify_sort thy (map (fn T => Sign.intern_class thy
berghofe@22530
   204
        ("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs);
berghofe@22530
   205
    val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
berghofe@22530
   206
    val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
berghofe@22530
   207
    val fsT = TFree (fs_ctxt_tyname, ind_sort);
berghofe@22530
   208
berghofe@24570
   209
    val inductive_forall_def' = Drule.instantiate'
berghofe@24570
   210
      [SOME (ctyp_of thy fsT)] [] inductive_forall_def;
berghofe@24570
   211
berghofe@22530
   212
    fun lift_pred' t (Free (s, T)) ts =
berghofe@22530
   213
      list_comb (Free (s, fsT --> T), t :: ts);
berghofe@22530
   214
    val lift_pred = lift_pred' (Bound 0);
berghofe@22530
   215
berghofe@24570
   216
    fun lift_prem (t as (f $ u)) =
berghofe@22530
   217
          let val (p, ts) = strip_comb t
berghofe@22530
   218
          in
berghofe@22530
   219
            if p mem ps then
berghofe@24570
   220
              Const (inductive_forall_name,
berghofe@24570
   221
                (fsT --> HOLogic.boolT) --> HOLogic.boolT) $
berghofe@24570
   222
                  Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
berghofe@22530
   223
            else lift_prem f $ lift_prem u
berghofe@22530
   224
          end
berghofe@22530
   225
      | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t)
berghofe@22530
   226
      | lift_prem t = t;
berghofe@22530
   227
berghofe@22530
   228
    fun mk_distinct [] = []
berghofe@22530
   229
      | mk_distinct ((x, T) :: xs) = List.mapPartial (fn (y, U) =>
berghofe@22530
   230
          if T = U then SOME (HOLogic.mk_Trueprop
berghofe@22530
   231
            (HOLogic.mk_not (HOLogic.eq_const T $ x $ y)))
berghofe@22530
   232
          else NONE) xs @ mk_distinct xs;
berghofe@22530
   233
berghofe@22530
   234
    fun mk_fresh (x, T) = HOLogic.mk_Trueprop
berghofe@25824
   235
      (NominalPackage.fresh_const T fsT $ x $ Bound 0);
berghofe@22530
   236
berghofe@22530
   237
    val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) =>
berghofe@22530
   238
      let
berghofe@22530
   239
        val params' = params @ [("y", fsT)];
berghofe@22530
   240
        val prem = Logic.list_implies
berghofe@22530
   241
          (map mk_fresh bvars @ mk_distinct bvars @
berghofe@22530
   242
           map (fn prem =>
berghofe@22530
   243
             if null (term_frees prem inter ps) then prem
berghofe@22530
   244
             else lift_prem prem) prems,
berghofe@22530
   245
           HOLogic.mk_Trueprop (lift_pred p ts));
berghofe@22530
   246
        val vs = map (Var o apfst (rpair 0)) (rename_wrt_term prem params')
berghofe@22530
   247
      in
berghofe@22530
   248
        (list_all (params', prem), (rev vs, subst_bounds (vs, prem)))
berghofe@22530
   249
      end) prems);
berghofe@22530
   250
berghofe@22530
   251
    val ind_vars =
berghofe@22530
   252
      (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~
berghofe@22530
   253
       map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
berghofe@22530
   254
    val ind_Ts = rev (map snd ind_vars);
berghofe@22530
   255
berghofe@22530
   256
    val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
berghofe@22530
   257
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
berghofe@22530
   258
        HOLogic.list_all (ind_vars, lift_pred p
berghofe@22530
   259
          (map (fold_rev (NominalPackage.mk_perm ind_Ts)
berghofe@22530
   260
            (map Bound (length atomTs downto 1))) ts)))) concls));
berghofe@22530
   261
berghofe@22530
   262
    val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
berghofe@22530
   263
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
berghofe@22530
   264
        lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls));
berghofe@22530
   265
berghofe@22530
   266
    val vc_compat = map (fn (params, bvars, prems, (p, ts)) =>
berghofe@22530
   267
      map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies
berghofe@24570
   268
          (List.mapPartial (fn prem =>
berghofe@24570
   269
             if null (ps inter term_frees prem) then SOME prem
berghofe@24570
   270
             else map_term (split_conj (K o I) names) prem prem) prems, q))))
berghofe@22530
   271
        (mk_distinct bvars @
berghofe@22530
   272
         maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
berghofe@25824
   273
           (NominalPackage.fresh_const U T $ u $ t)) bvars)
berghofe@22530
   274
             (ts ~~ binder_types (fastype_of p)))) prems;
berghofe@22530
   275
wenzelm@26343
   276
    val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp";
wenzelm@26343
   277
    val pt2_atoms = map (fn aT => PureThy.get_thm thy
wenzelm@26337
   278
      ("pt_" ^ Sign.base_name (fst (dest_Type aT)) ^ "2")) atomTs;
berghofe@27352
   279
    val eqvt_ss = Simplifier.theory_context thy HOL_basic_ss
berghofe@27352
   280
      addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms)
berghofe@27352
   281
      addsimprocs [mk_perm_bool_simproc ["Fun.id"],
berghofe@27352
   282
        NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun];
wenzelm@26343
   283
    val fresh_bij = PureThy.get_thms thy "fresh_bij";
wenzelm@26343
   284
    val perm_bij = PureThy.get_thms thy "perm_bij";
wenzelm@26343
   285
    val fs_atoms = map (fn aT => PureThy.get_thm thy
wenzelm@26337
   286
      ("fs_" ^ Sign.base_name (fst (dest_Type aT)) ^ "1")) atomTs;
wenzelm@26343
   287
    val exists_fresh' = PureThy.get_thms thy "exists_fresh'";
wenzelm@26343
   288
    val fresh_atm = PureThy.get_thms thy "fresh_atm";
wenzelm@26343
   289
    val calc_atm = PureThy.get_thms thy "calc_atm";
wenzelm@26343
   290
    val perm_fresh_fresh = PureThy.get_thms thy "perm_fresh_fresh";
berghofe@22530
   291
berghofe@22530
   292
    fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) =
berghofe@22530
   293
      let
berghofe@22530
   294
        (** protect terms to avoid that supp_prod interferes with   **)
berghofe@22530
   295
        (** pairs used in introduction rules of inductive predicate **)
berghofe@22530
   296
        fun protect t =
berghofe@22530
   297
          let val T = fastype_of t in Const ("Fun.id", T --> T) $ t end;
berghofe@22530
   298
        val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1);
berghofe@22530
   299
        val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
berghofe@22530
   300
            (HOLogic.exists_const T $ Abs ("x", T,
berghofe@25824
   301
              NominalPackage.fresh_const T (fastype_of p) $
berghofe@22530
   302
                Bound 0 $ p)))
berghofe@22530
   303
          (fn _ => EVERY
berghofe@22530
   304
            [resolve_tac exists_fresh' 1,
berghofe@22530
   305
             simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
berghofe@22530
   306
        val (([cx], ths), ctxt') = Obtain.result
berghofe@22530
   307
          (fn _ => EVERY
berghofe@22530
   308
            [etac exE 1,
berghofe@22530
   309
             full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
haftmann@26359
   310
             full_simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1,
berghofe@22530
   311
             REPEAT (etac conjE 1)])
berghofe@22530
   312
          [ex] ctxt
berghofe@22530
   313
      in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
berghofe@22530
   314
berghofe@25824
   315
    fun mk_ind_proof thy thss =
wenzelm@26711
   316
      Goal.prove_global thy [] prems' concl' (fn {prems = ihyps, context = ctxt} =>
berghofe@22530
   317
        let val th = Goal.prove ctxt [] [] concl (fn {context, ...} =>
berghofe@22530
   318
          rtac raw_induct 1 THEN
berghofe@22530
   319
          EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) =>
berghofe@22530
   320
            [REPEAT (rtac allI 1), simp_tac eqvt_ss 1,
berghofe@22530
   321
             SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
berghofe@22530
   322
               let
berghofe@22530
   323
                 val (params', (pis, z)) =
berghofe@22530
   324
                   chop (length params - length atomTs - 1) (map term_of params) ||>
berghofe@22530
   325
                   split_last;
berghofe@22530
   326
                 val bvars' = map
berghofe@22530
   327
                   (fn (Bound i, T) => (nth params' (length params' - i), T)
berghofe@22530
   328
                     | (t, T) => (t, T)) bvars;
berghofe@22530
   329
                 val pi_bvars = map (fn (t, _) =>
berghofe@22530
   330
                   fold_rev (NominalPackage.mk_perm []) pis t) bvars';
berghofe@22530
   331
                 val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
berghofe@22530
   332
                 val (freshs1, freshs2, ctxt'') = fold
berghofe@22530
   333
                   (obtain_fresh_name (ts @ pi_bvars))
berghofe@22530
   334
                   (map snd bvars') ([], [], ctxt');
berghofe@22530
   335
                 val freshs2' = NominalPackage.mk_not_sym freshs2;
berghofe@22530
   336
                 val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1);
berghofe@24570
   337
                 fun concat_perm pi1 pi2 =
berghofe@24570
   338
                   let val T = fastype_of pi1
berghofe@24570
   339
                   in if T = fastype_of pi2 then
berghofe@24570
   340
                       Const ("List.append", T --> T --> T) $ pi1 $ pi2
berghofe@24570
   341
                     else pi2
berghofe@24570
   342
                   end;
berghofe@24570
   343
                 val pis'' = fold (concat_perm #> map) pis' pis;
berghofe@22530
   344
                 val env = Pattern.first_order_match thy (ihypt, prop_of ihyp)
berghofe@22530
   345
                   (Vartab.empty, Vartab.empty);
berghofe@22530
   346
                 val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy))
berghofe@22530
   347
                   (map (Envir.subst_vars env) vs ~~
berghofe@22530
   348
                    map (fold_rev (NominalPackage.mk_perm [])
berghofe@22530
   349
                      (rev pis' @ pis)) params' @ [z])) ihyp;
berghofe@24570
   350
                 fun mk_pi th =
haftmann@26359
   351
                   Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
berghofe@24570
   352
                       addsimprocs [NominalPackage.perm_simproc])
berghofe@24570
   353
                     (Simplifier.simplify eqvt_ss
berghofe@25824
   354
                       (fold_rev (mk_perm_bool o cterm_of thy)
berghofe@25824
   355
                         (rev pis' @ pis) th));
berghofe@24570
   356
                 val (gprems1, gprems2) = split_list
berghofe@24570
   357
                   (map (fn (th, t) =>
berghofe@24570
   358
                      if null (term_frees t inter ps) then (SOME th, mk_pi th)
berghofe@24570
   359
                      else
berghofe@24570
   360
                        (map_thm ctxt (split_conj (K o I) names)
berghofe@24570
   361
                           (etac conjunct1 1) monos NONE th,
berghofe@24570
   362
                         mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis''))
berghofe@24570
   363
                           (inst_conj_all_tac (length pis'')) monos (SOME t) th))))
berghofe@24570
   364
                      (gprems ~~ oprems)) |>> List.mapPartial I;
berghofe@22530
   365
                 val vc_compat_ths' = map (fn th =>
berghofe@22530
   366
                   let
berghofe@25824
   367
                     val th' = first_order_mrs gprems1 th;
berghofe@22530
   368
                     val (bop, lhs, rhs) = (case concl_of th' of
berghofe@22530
   369
                         _ $ (fresh $ lhs $ rhs) =>
berghofe@22530
   370
                           (fn t => fn u => fresh $ t $ u, lhs, rhs)
berghofe@22530
   371
                       | _ $ (_ $ (_ $ lhs $ rhs)) =>
berghofe@22530
   372
                           (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs));
berghofe@22530
   373
                     val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop
berghofe@22530
   374
                         (bop (fold_rev (NominalPackage.mk_perm []) pis lhs)
berghofe@22530
   375
                            (fold_rev (NominalPackage.mk_perm []) pis rhs)))
berghofe@22530
   376
                       (fn _ => simp_tac (HOL_basic_ss addsimps
berghofe@22530
   377
                          (fresh_bij @ perm_bij)) 1 THEN rtac th' 1)
berghofe@22530
   378
                   in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end)
berghofe@22530
   379
                     vc_compat_ths;
berghofe@22530
   380
                 val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths';
berghofe@22530
   381
                 (** Since calc_atm simplifies (pi :: 'a prm) o (x :: 'b) to x **)
berghofe@22530
   382
                 (** we have to pre-simplify the rewrite rules                 **)
berghofe@22530
   383
                 val calc_atm_ss = HOL_ss addsimps calc_atm @
berghofe@22530
   384
                    map (Simplifier.simplify (HOL_ss addsimps calc_atm))
berghofe@22530
   385
                      (vc_compat_ths'' @ freshs2');
berghofe@22530
   386
                 val th = Goal.prove ctxt'' [] []
berghofe@22530
   387
                   (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
berghofe@22530
   388
                     map (fold (NominalPackage.mk_perm []) pis') (tl ts))))
berghofe@22530
   389
                   (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1,
berghofe@22530
   390
                     REPEAT_DETERM_N (nprems_of ihyp - length gprems)
berghofe@22530
   391
                       (simp_tac calc_atm_ss 1),
berghofe@22530
   392
                     REPEAT_DETERM_N (length gprems)
berghofe@24570
   393
                       (simp_tac (HOL_ss
berghofe@24570
   394
                          addsimps inductive_forall_def' :: gprems2
berghofe@22530
   395
                          addsimprocs [NominalPackage.perm_simproc]) 1)]));
berghofe@22530
   396
                 val final = Goal.prove ctxt'' [] [] (term_of concl)
berghofe@22530
   397
                   (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
berghofe@22530
   398
                     addsimps vc_compat_ths'' @ freshs2' @
berghofe@22530
   399
                       perm_fresh_fresh @ fresh_atm) 1);
berghofe@22530
   400
                 val final' = ProofContext.export ctxt'' ctxt' [final];
berghofe@22530
   401
               in resolve_tac final' 1 end) context 1])
berghofe@22530
   402
                 (prems ~~ thss ~~ ihyps ~~ prems'')))
berghofe@22530
   403
        in
berghofe@22530
   404
          cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN
berghofe@22530
   405
          REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN
wenzelm@27228
   406
            etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN
berghofe@22530
   407
            asm_full_simp_tac (simpset_of thy) 1)
wenzelm@26711
   408
        end);
berghofe@22530
   409
berghofe@25824
   410
    (** strong case analysis rule **)
berghofe@25824
   411
berghofe@25824
   412
    val cases_prems = map (fn ((name, avoids), rule) =>
berghofe@25824
   413
      let
berghofe@25824
   414
        val prem :: prems = Logic.strip_imp_prems rule;
berghofe@25824
   415
        val concl = Logic.strip_imp_concl rule;
berghofe@25824
   416
        val used = add_term_free_names (rule, [])
berghofe@25824
   417
      in
berghofe@25824
   418
        (prem,
berghofe@25824
   419
         List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params),
berghofe@25824
   420
         concl,
berghofe@25824
   421
         fst (fold_map (fn (prem, (_, avoid)) => fn used =>
berghofe@25824
   422
           let
berghofe@25824
   423
             val prems = Logic.strip_assums_hyp prem;
berghofe@25824
   424
             val params = Logic.strip_params prem;
berghofe@25824
   425
             val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid;
berghofe@25824
   426
             fun mk_subst (p as (s, T)) (i, j, used, ps, qs, is, ts) =
berghofe@25824
   427
               if member (op = o apsnd fst) bnds (Bound i) then
berghofe@25824
   428
                 let
berghofe@25824
   429
                   val s' = Name.variant used s;
berghofe@25824
   430
                   val t = Free (s', T)
berghofe@25824
   431
                 in (i + 1, j, s' :: used, ps, (t, T) :: qs, i :: is, t :: ts) end
berghofe@25824
   432
               else (i + 1, j + 1, used, p :: ps, qs, is, Bound j :: ts);
berghofe@25824
   433
             val (_, _, used', ps, qs, is, ts) = fold_rev mk_subst params
berghofe@25824
   434
               (0, 0, used, [], [], [], [])
berghofe@25824
   435
           in
berghofe@25824
   436
             ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), used')
berghofe@25824
   437
           end) (prems ~~ avoids) used))
berghofe@25824
   438
      end)
berghofe@25824
   439
        (InductivePackage.partition_rules' raw_induct (intrs ~~ avoids') ~~
berghofe@25824
   440
         elims');
berghofe@25824
   441
berghofe@25824
   442
    val cases_prems' =
berghofe@25824
   443
      map (fn (prem, args, concl, prems) =>
berghofe@25824
   444
        let
berghofe@25824
   445
          fun mk_prem (ps, [], _, prems) =
berghofe@25824
   446
                list_all (ps, Logic.list_implies (prems, concl))
berghofe@25824
   447
            | mk_prem (ps, qs, _, prems) =
berghofe@25824
   448
                list_all (ps, Logic.mk_implies
berghofe@25824
   449
                  (Logic.list_implies
berghofe@25824
   450
                    (mk_distinct qs @
berghofe@25824
   451
                     maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop
berghofe@25824
   452
                      (NominalPackage.fresh_const T (fastype_of u) $ t $ u))
berghofe@25824
   453
                        args) qs,
berghofe@25824
   454
                     HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
berghofe@25824
   455
                       (map HOLogic.dest_Trueprop prems))),
berghofe@25824
   456
                   concl))
berghofe@25824
   457
          in map mk_prem prems end) cases_prems;
berghofe@25824
   458
berghofe@27352
   459
    val cases_eqvt_ss = Simplifier.theory_context thy HOL_ss
berghofe@27352
   460
      addsimps eqvt_thms @ fresh_atm @ perm_pi_simp delsplits [split_if]
berghofe@27352
   461
      addsimprocs [NominalPermeq.perm_simproc_app,
berghofe@27352
   462
        NominalPermeq.perm_simproc_fun];
berghofe@27352
   463
berghofe@27352
   464
    val simp_fresh_atm = map
berghofe@27352
   465
      (Simplifier.simplify (HOL_basic_ss addsimps fresh_atm));
berghofe@25824
   466
berghofe@25824
   467
    fun mk_cases_proof thy ((((name, thss), elim), (prem, args, concl, prems)),
berghofe@25824
   468
        prems') =
wenzelm@26711
   469
      (name, Goal.prove_global thy [] (prem :: prems') concl
wenzelm@26711
   470
        (fn {prems = hyp :: hyps, context = ctxt1} =>
berghofe@25824
   471
        EVERY (rtac (hyp RS elim) 1 ::
berghofe@25824
   472
          map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) =>
berghofe@25824
   473
            SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} =>
berghofe@25824
   474
              if null qs then
berghofe@25824
   475
                rtac (first_order_mrs case_hyps case_hyp) 1
berghofe@25824
   476
              else
berghofe@25824
   477
                let
berghofe@25824
   478
                  val params' = map (term_of o nth (rev params)) is;
berghofe@25824
   479
                  val tab = params' ~~ map fst qs;
berghofe@25824
   480
                  val (hyps1, hyps2) = chop (length args) case_hyps;
berghofe@25824
   481
                  (* turns a = t and [x1 # t, ..., xn # t] *)
berghofe@25824
   482
                  (* into [x1 # a, ..., xn # a]            *)
berghofe@25824
   483
                  fun inst_fresh th' ths =
berghofe@25824
   484
                    let val (ths1, ths2) = chop (length qs) ths
berghofe@25824
   485
                    in
berghofe@25824
   486
                      (map (fn th =>
berghofe@25824
   487
                         let
berghofe@25824
   488
                           val (cf, ct) =
berghofe@25824
   489
                             Thm.dest_comb (Thm.dest_arg (cprop_of th));
berghofe@25824
   490
                           val arg_cong' = Drule.instantiate'
berghofe@25824
   491
                             [SOME (ctyp_of_term ct)]
berghofe@25824
   492
                             [NONE, SOME ct, SOME cf] (arg_cong RS iffD2);
berghofe@25824
   493
                           val inst = Thm.first_order_match (ct,
berghofe@25824
   494
                             Thm.dest_arg (Thm.dest_arg (cprop_of th')))
berghofe@25824
   495
                         in [th', th] MRS Thm.instantiate inst arg_cong'
berghofe@25824
   496
                         end) ths1,
berghofe@25824
   497
                       ths2)
berghofe@25824
   498
                    end;
berghofe@25824
   499
                  val (vc_compat_ths1, vc_compat_ths2) =
berghofe@25824
   500
                    chop (length vc_compat_ths - length args * length qs)
berghofe@25824
   501
                      (map (first_order_mrs hyps2) vc_compat_ths);
berghofe@25824
   502
                  val vc_compat_ths' =
berghofe@25824
   503
                    NominalPackage.mk_not_sym vc_compat_ths1 @
berghofe@25824
   504
                    flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2));
berghofe@25824
   505
                  val (freshs1, freshs2, ctxt3) = fold
berghofe@25824
   506
                    (obtain_fresh_name (args @ map fst qs @ params'))
berghofe@25824
   507
                    (map snd qs) ([], [], ctxt2);
berghofe@25824
   508
                  val freshs2' = NominalPackage.mk_not_sym freshs2;
berghofe@25824
   509
                  val pis = map (NominalPackage.perm_of_pair)
berghofe@25824
   510
                    ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1));
berghofe@25824
   511
                  val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis);
berghofe@25824
   512
                  val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
berghofe@25824
   513
                     (fn x as Free _ =>
berghofe@25824
   514
                           if x mem args then x
berghofe@25824
   515
                           else (case AList.lookup op = tab x of
berghofe@25824
   516
                             SOME y => y
berghofe@25824
   517
                           | NONE => fold_rev (NominalPackage.mk_perm []) pis x)
berghofe@25824
   518
                       | x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps));
berghofe@25824
   519
                  val inst = Thm.first_order_match (Thm.dest_arg
berghofe@25824
   520
                    (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
berghofe@25824
   521
                  val th = Goal.prove ctxt3 [] [] (term_of concl)
berghofe@25824
   522
                    (fn {context = ctxt4, ...} =>
berghofe@25824
   523
                       rtac (Thm.instantiate inst case_hyp) 1 THEN
berghofe@25824
   524
                       SUBPROOF (fn {prems = fresh_hyps, ...} =>
berghofe@25824
   525
                         let
berghofe@25824
   526
                           val fresh_hyps' = NominalPackage.mk_not_sym fresh_hyps;
berghofe@27352
   527
                           val case_ss = cases_eqvt_ss addsimps freshs2' @
berghofe@27352
   528
                             simp_fresh_atm (vc_compat_ths' @ fresh_hyps');
berghofe@25824
   529
                           val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh;
berghofe@27352
   530
                           val calc_atm_ss = case_ss addsimps calc_atm;
berghofe@25824
   531
                           val hyps1' = map
berghofe@25824
   532
                             (mk_pis #> Simplifier.simplify fresh_fresh_ss) hyps1;
berghofe@25824
   533
                           val hyps2' = map
berghofe@25824
   534
                             (mk_pis #> Simplifier.simplify case_ss) hyps2;
berghofe@27352
   535
                           (* calc_atm must be applied last, since *)
berghofe@27352
   536
                           (* it may interfere with other rules    *)
berghofe@27352
   537
                           val case_hyps' = map
berghofe@27352
   538
                             (Simplifier.simplify calc_atm_ss) (hyps1' @ hyps2')
berghofe@25824
   539
                         in
berghofe@27352
   540
                           simp_tac calc_atm_ss 1 THEN
berghofe@25824
   541
                           REPEAT_DETERM (TRY (rtac conjI 1) THEN
berghofe@25824
   542
                             resolve_tac case_hyps' 1)
berghofe@25824
   543
                         end) ctxt4 1)
berghofe@25824
   544
                  val final = ProofContext.export ctxt3 ctxt2 [th]
berghofe@25824
   545
                in resolve_tac final 1 end) ctxt1 1)
berghofe@25824
   546
                  (thss ~~ hyps ~~ prems))))
berghofe@25824
   547
berghofe@22530
   548
  in
berghofe@22530
   549
    thy |>
berghofe@22530
   550
    ProofContext.init |>
berghofe@22530
   551
    Proof.theorem_i NONE (fn thss => ProofContext.theory (fn thy =>
berghofe@22530
   552
      let
berghofe@22530
   553
        val ctxt = ProofContext.init thy;
berghofe@22530
   554
        val rec_name = space_implode "_" (map Sign.base_name names);
berghofe@22530
   555
        val ind_case_names = RuleCases.case_names induct_cases;
berghofe@25824
   556
        val induct_cases' = InductivePackage.partition_rules' raw_induct
berghofe@25824
   557
          (intrs ~~ induct_cases); 
berghofe@25824
   558
        val thss' = map (map atomize_intr) thss;
berghofe@25824
   559
        val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss');
berghofe@24570
   560
        val strong_raw_induct =
berghofe@25824
   561
          mk_ind_proof thy thss' |> InductivePackage.rulify;
berghofe@25824
   562
        val strong_cases = map (mk_cases_proof thy ##> InductivePackage.rulify)
berghofe@25824
   563
          (thsss ~~ elims ~~ cases_prems ~~ cases_prems');
berghofe@22530
   564
        val strong_induct =
berghofe@22530
   565
          if length names > 1 then
berghofe@22530
   566
            (strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
berghofe@22530
   567
          else (strong_raw_induct RSN (2, rev_mp),
berghofe@22530
   568
            [ind_case_names, RuleCases.consumes 1]);
berghofe@22530
   569
        val ([strong_induct'], thy') = thy |>
wenzelm@24712
   570
          Sign.add_path rec_name |>
berghofe@22530
   571
          PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)];
berghofe@22530
   572
        val strong_inducts =
berghofe@22530
   573
          ProjectRule.projects ctxt (1 upto length names) strong_induct'
berghofe@22530
   574
      in
berghofe@22530
   575
        thy' |>
berghofe@22530
   576
        PureThy.add_thmss [(("strong_inducts", strong_inducts),
berghofe@22530
   577
          [ind_case_names, RuleCases.consumes 1])] |> snd |>
berghofe@25824
   578
        Sign.parent_path |>
berghofe@25824
   579
        fold (fn ((name, elim), (_, cases)) =>
berghofe@25824
   580
          Sign.add_path (Sign.base_name name) #>
berghofe@25824
   581
          PureThy.add_thms [(("strong_cases", elim),
berghofe@25824
   582
            [RuleCases.case_names (map snd cases),
berghofe@25824
   583
             RuleCases.consumes 1])] #> snd #>
berghofe@25824
   584
          Sign.parent_path) (strong_cases ~~ induct_cases')
berghofe@22530
   585
      end))
berghofe@24570
   586
      (map (map (rulify_term thy #> rpair [])) vc_compat)
berghofe@22530
   587
  end;
berghofe@22530
   588
berghofe@22730
   589
fun prove_eqvt s xatoms thy =
berghofe@22530
   590
  let
berghofe@22530
   591
    val ctxt = ProofContext.init thy;
berghofe@22788
   592
    val ({names, ...}, {raw_induct, intrs, elims, ...}) =
berghofe@22730
   593
      InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
wenzelm@24832
   594
    val raw_induct = atomize_induct ctxt raw_induct;
wenzelm@24832
   595
    val elims = map (atomize_induct ctxt) elims;
berghofe@24570
   596
    val intrs = map atomize_intr intrs;
berghofe@24570
   597
    val monos = InductivePackage.get_monos ctxt;
berghofe@22788
   598
    val intrs' = InductivePackage.unpartition_rules intrs
berghofe@22788
   599
      (map (fn (((s, ths), (_, k)), th) =>
berghofe@22788
   600
           (s, ths ~~ InductivePackage.infer_intro_vars th k ths))
berghofe@22788
   601
         (InductivePackage.partition_rules raw_induct intrs ~~
berghofe@22788
   602
          InductivePackage.arities_of raw_induct ~~ elims));
berghofe@22730
   603
    val atoms' = NominalAtoms.atoms_of thy;
berghofe@22730
   604
    val atoms =
berghofe@22730
   605
      if null xatoms then atoms' else
berghofe@22730
   606
      let val atoms = map (Sign.intern_type thy) xatoms
berghofe@22730
   607
      in
berghofe@22730
   608
        (case duplicates op = atoms of
berghofe@22730
   609
             [] => ()
berghofe@22730
   610
           | xs => error ("Duplicate atoms: " ^ commas xs);
berghofe@22730
   611
         case atoms \\ atoms' of
berghofe@22730
   612
             [] => ()
berghofe@22730
   613
           | xs => error ("No such atoms: " ^ commas xs);
berghofe@22730
   614
         atoms)
berghofe@22730
   615
      end;
wenzelm@26343
   616
    val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp";
berghofe@26364
   617
    val eqvt_ss = Simplifier.theory_context thy HOL_basic_ss addsimps
urbanc@24571
   618
      (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs
berghofe@26364
   619
      [mk_perm_bool_simproc names,
berghofe@26364
   620
       NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun];
berghofe@22313
   621
    val t = Logic.unvarify (concl_of raw_induct);
berghofe@22313
   622
    val pi = Name.variant (add_term_names (t, [])) "pi";
berghofe@22313
   623
    val ps = map (fst o HOLogic.dest_imp)
berghofe@22313
   624
      (HOLogic.dest_conj (HOLogic.dest_Trueprop t));
berghofe@25824
   625
    fun eqvt_tac pi (intr, vs) st =
berghofe@22544
   626
      let
berghofe@22544
   627
        fun eqvt_err s = error
berghofe@22544
   628
          ("Could not prove equivariance for introduction rule\n" ^
wenzelm@26939
   629
           Syntax.string_of_term_global (theory_of_thm intr)
berghofe@22544
   630
             (Logic.unvarify (prop_of intr)) ^ "\n" ^ s);
berghofe@22788
   631
        val res = SUBPROOF (fn {prems, params, ...} =>
berghofe@22788
   632
          let
berghofe@24570
   633
            val prems' = map (fn th => the_default th (map_thm ctxt
berghofe@24570
   634
              (split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems;
berghofe@25824
   635
            val prems'' = map (fn th => Simplifier.simplify eqvt_ss
berghofe@25824
   636
              (mk_perm_bool (cterm_of thy pi) th)) prems';
berghofe@22788
   637
            val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
berghofe@22788
   638
               map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params)
berghofe@22788
   639
               intr
berghofe@24570
   640
          in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
berghofe@22544
   641
          end) ctxt 1 st
berghofe@22544
   642
      in
berghofe@22544
   643
        case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of
berghofe@22544
   644
          NONE => eqvt_err ("Rule does not match goal\n" ^
wenzelm@26939
   645
            Syntax.string_of_term_global (theory_of_thm st) (hd (prems_of st)))
berghofe@22544
   646
        | SOME (th, _) => Seq.single th
berghofe@22544
   647
      end;
berghofe@22313
   648
    val thss = map (fn atom =>
berghofe@25824
   649
      let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, [])))
berghofe@22530
   650
      in map (fn th => zero_var_indexes (th RS mp))
berghofe@22313
   651
        (DatatypeAux.split_conj_thm (Goal.prove_global thy [] []
berghofe@22313
   652
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p =>
berghofe@22313
   653
            HOLogic.mk_imp (p, list_comb
berghofe@22313
   654
             (apsnd (map (NominalPackage.mk_perm [] pi')) (strip_comb p)))) ps)))
berghofe@22788
   655
          (fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs =>
berghofe@22788
   656
              full_simp_tac eqvt_ss 1 THEN
berghofe@25824
   657
              eqvt_tac pi' intr_vs) intrs'))))
berghofe@22544
   658
      end) atoms
berghofe@22544
   659
  in
berghofe@22544
   660
    fold (fn (name, ths) =>
wenzelm@24712
   661
      Sign.add_path (Sign.base_name name) #>
berghofe@22544
   662
      PureThy.add_thmss [(("eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #>
wenzelm@24712
   663
      Sign.parent_path) (names ~~ transp thss) thy
berghofe@22544
   664
  end;
berghofe@22313
   665
berghofe@22313
   666
berghofe@22313
   667
(* outer syntax *)
berghofe@22313
   668
berghofe@22313
   669
local structure P = OuterParse and K = OuterKeyword in
berghofe@22313
   670
wenzelm@27353
   671
val _ = OuterKeyword.keyword "avoids";
wenzelm@24867
   672
wenzelm@24867
   673
val _ =
berghofe@22313
   674
  OuterSyntax.command "nominal_inductive"
berghofe@22530
   675
    "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
berghofe@22530
   676
    (P.name -- Scan.optional (P.$$$ "avoids" |-- P.and_list1 (P.name --
berghofe@22530
   677
      (P.$$$ ":" |-- Scan.repeat1 P.name))) [] >> (fn (name, avoids) =>
berghofe@22730
   678
        Toplevel.print o Toplevel.theory_to_proof (prove_strong_ind name avoids)));
berghofe@22313
   679
wenzelm@24867
   680
val _ =
berghofe@22530
   681
  OuterSyntax.command "equivariance"
berghofe@22530
   682
    "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
berghofe@22730
   683
    (P.name -- Scan.optional (P.$$$ "[" |-- P.list1 P.name --| P.$$$ "]") [] >>
berghofe@22730
   684
      (fn (name, atoms) => Toplevel.theory (prove_eqvt name atoms)));
berghofe@22530
   685
berghofe@22313
   686
end;
berghofe@22313
   687
berghofe@22313
   688
end