src/HOL/Nominal/nominal_inductive2.ML
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 82643 f1c14af17591
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_inductive2.ML
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     3
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     4
Infrastructure for proving equivariance and strong induction theorems
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     5
for inductive predicates involving nominal datatypes.
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     6
Experimental version that allows to avoid lists of atoms.
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     7
*)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     8
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
     9
signature NOMINAL_INDUCTIVE2 =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    10
sig
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
    11
  val prove_strong_ind: string -> string option -> (string * string list) list ->
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
    12
    local_theory -> Proof.state
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    13
end
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    14
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    15
structure NominalInductive2 : NOMINAL_INDUCTIVE2 =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    16
struct
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    17
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59936
diff changeset
    18
val inductive_forall_def = @{thm HOL.induct_forall_def};
33772
b6a1feca2ac2 use of thm-antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 33671
diff changeset
    19
val inductive_atomize = @{thms induct_atomize};
b6a1feca2ac2 use of thm-antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 33671
diff changeset
    20
val inductive_rulify = @{thms induct_rulify};
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    21
82643
f1c14af17591 prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents: 81946
diff changeset
    22
fun rulify_term thy = Simplifier.rewrite_term thy inductive_rulify [];
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    23
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    24
fun atomize_conv ctxt =
82643
f1c14af17591 prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents: 81946
diff changeset
    25
  Simplifier.rewrite_cterm (true, false, false) (K (K NONE))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    26
    (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    27
fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt));
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    28
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
    29
  (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt));
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    30
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    31
fun fresh_postprocess ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    32
  Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps
30108
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
    33
    [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim},
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
    34
     @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]);
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
    35
44929
haftmann
parents: 44868
diff changeset
    36
fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []);
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
    37
44689
f247fc952f31 tuned specifications
haftmann
parents: 44241
diff changeset
    38
val perm_bool = mk_meta_eq @{thm perm_bool_def};
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38795
diff changeset
    39
val perm_boolI = @{thm perm_boolI};
81946
ee680c69de38 misc tuning: prefer specific variants of Thm.dest_comb;
wenzelm
parents: 81541
diff changeset
    40
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (Thm.dest_arg
ee680c69de38 misc tuning: prefer specific variants of Thm.dest_comb;
wenzelm
parents: 81541
diff changeset
    41
  (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)));
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    42
60787
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
    43
fun mk_perm_bool ctxt pi th =
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
    44
  th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    45
60359
cb8828b859a1 clarified context;
wenzelm
parents: 59940
diff changeset
    46
fun mk_perm_bool_simproc names =
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    47
  Simplifier.make_simproc \<^context>
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    48
   {name = "perm_bool",
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80703
diff changeset
    49
    kind = Simproc,
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    50
    lhss = [\<^term>\<open>perm pi x\<close>],
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    51
    proc = fn _ => fn _ => fn ct =>
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    52
      (case Thm.term_of ct of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
    53
        Const (\<^const_name>\<open>Nominal.perm\<close>, _) $ _ $ t =>
80636
4041e7c8059d tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents: 80634
diff changeset
    54
          if member (op =) names (the_default "" (try (dest_Const_name o head_of) t))
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    55
          then SOME perm_bool else NONE
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    56
       | _ => NONE),
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    57
    identifier = []};
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    58
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    59
fun transp ([] :: _) = []
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    60
  | transp xs = map hd xs :: transp (map tl xs);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    61
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    62
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    63
      (Const (s, T), ts) => (case strip_type T of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    64
        (Ts, Type (tname, _)) =>
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
    65
          (case NominalDatatype.get_nominal_datatype thy tname of
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    66
             NONE => fold (add_binders thy i) ts bs
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    67
           | SOME {descr, index, ...} => (case AList.lookup op =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    68
                 (#3 (the (AList.lookup op = descr index))) s of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    69
               NONE => fold (add_binders thy i) ts bs
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    70
             | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    71
                 let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs'
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    72
                 in (add_binders thy i u
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    73
                   (fold (fn (u, T) =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    74
                      if exists (fn j => j < i) (loose_bnos u) then I
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    75
                      else AList.map_default op = (T, [])
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    76
                        (insert op aconv (incr_boundvars (~i) u)))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    77
                          cargs1 bs'), cargs2)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    78
                 end) cargs (bs, ts ~~ Ts))))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    79
      | _ => fold (add_binders thy i) ts bs)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    80
    | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    81
  | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    82
  | add_binders thy i _ bs = bs;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    83
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
    84
fun split_conj f names (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ = (case head_of p of
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    85
      Const (name, _) =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36428
diff changeset
    86
        if member (op =) names name then SOME (f p q) else NONE
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    87
    | _ => NONE)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    88
  | split_conj _ _ _ _ = NONE;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    89
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    90
fun strip_all [] t = t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
    91
  | strip_all (_ :: xs) (Const (\<^const_name>\<open>All\<close>, _) $ Abs (s, T, t)) = strip_all xs t;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    92
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    93
(*********************************************************************)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    94
(* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    95
(* or    ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)            *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    96
(* to    R ... & id (ALL z. P z (pi_1 o ... o pi_n o t))             *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    97
(* or    id (ALL z. P z (pi_1 o ... o pi_n o t))                     *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    98
(*                                                                   *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
    99
(* where "id" protects the subformula from simplification            *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   100
(*********************************************************************)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   101
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   102
fun inst_conj_all names ps pis (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ =
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   103
      (case head_of p of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   104
         Const (name, _) =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36428
diff changeset
   105
           if member (op =) names name then SOME (HOLogic.mk_conj (p,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   106
             Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   107
               (subst_bounds (pis, strip_all pis q))))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   108
           else NONE
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   109
       | _ => NONE)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   110
  | inst_conj_all names ps pis t u =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   111
      if member (op aconv) ps (head_of u) then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   112
        SOME (Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   113
          (subst_bounds (pis, strip_all pis t)))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   114
      else NONE
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   115
  | inst_conj_all _ _ _ _ _ = NONE;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   116
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   117
fun inst_conj_all_tac ctxt k = EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   118
  [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   119
   REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   120
   simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1];
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   121
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   122
fun map_term f t u = (case f t u of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   123
      NONE => map_term' f t u | x => x)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   124
and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   125
      (NONE, NONE) => NONE
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   126
    | (SOME t'', NONE) => SOME (t'' $ u)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   127
    | (NONE, SOME u'') => SOME (t $ u'')
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   128
    | (SOME t'', SOME u'') => SOME (t'' $ u''))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   129
  | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   130
      NONE => NONE
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   131
    | SOME t'' => SOME (Abs (s, T, t'')))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   132
  | map_term' _ _ _ = NONE;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   133
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   134
(*********************************************************************)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   135
(*         Prove  F[f t]  from  F[t],  where F is monotone           *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   136
(*********************************************************************)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   137
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   138
fun map_thm ctxt f tac monos opt th =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   139
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   140
    val prop = Thm.prop_of th;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   141
    fun prove t =
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   142
      Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   143
        EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   144
          REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)),
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   145
          REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))])
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   146
  in Option.map prove (map_term f prop (the_default prop opt)) end;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   147
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   148
fun abs_params params t =
81541
5335b1ca6233 more elementary operation Term.variant_bounds: only for bounds vs. frees, no consts, no tfrees;
wenzelm
parents: 81516
diff changeset
   149
  let val vs =  map (Var o apfst (rpair 0)) (Term.variant_bounds t params)
81516
31b05aef022d eliminate historic clone (see also 550e36c6a2d1);
wenzelm
parents: 80709
diff changeset
   150
  in (Logic.list_all (params, t), (vs, subst_bounds (rev vs, t))) end;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   151
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   152
fun inst_params thy (vs, p) th cts =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   153
  let val env = Pattern.first_order_match thy (p, Thm.prop_of th)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   154
    (Vartab.empty, Vartab.empty)
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   155
  in
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   156
    Thm.instantiate (TVars.empty, Vars.make (map (dest_Var o Envir.subst_term env) vs ~~ cts)) th
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   157
  end;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   158
74523
wenzelm
parents: 74522
diff changeset
   159
fun prove_strong_ind s alt_name avoids lthy =
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   160
  let
74523
wenzelm
parents: 74522
diff changeset
   161
    val thy = Proof_Context.theory_of lthy;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   162
    val ({names, ...}, {raw_induct, intrs, elims, ...}) =
74523
wenzelm
parents: 74522
diff changeset
   163
      Inductive.the_inductive_global lthy (Sign.intern_const thy s);
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   164
    val ind_params = Inductive.params_of raw_induct;
74523
wenzelm
parents: 74522
diff changeset
   165
    val raw_induct = atomize_induct lthy raw_induct;
wenzelm
parents: 74522
diff changeset
   166
    val elims = map (atomize_induct lthy) elims;
wenzelm
parents: 74522
diff changeset
   167
    val monos = Inductive.get_monos lthy;
wenzelm
parents: 74522
diff changeset
   168
    val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy;
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   169
    val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   170
        [] => ()
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   171
      | xs => error ("Missing equivariance theorem for predicate(s): " ^
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   172
          commas_quote xs));
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 43326
diff changeset
   173
    val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the
74523
wenzelm
parents: 74522
diff changeset
   174
      (Induct.lookup_inductP lthy (hd names)))));
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   175
    val induct_cases' = if null induct_cases then replicate (length intrs) ""
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   176
      else induct_cases;
74523
wenzelm
parents: 74522
diff changeset
   177
    val (raw_induct', ctxt') = lthy
70326
wenzelm
parents: 69597
diff changeset
   178
      |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   179
    val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   180
      HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   181
    val ps = map (fst o snd) concls;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   182
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58112
diff changeset
   183
    val _ = (case duplicates (op = o apply2 fst) avoids of
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   184
        [] => ()
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   185
      | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   186
    val _ = (case subtract (op =) induct_cases (map fst avoids) of
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   187
        [] => ()
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   188
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   189
    fun mk_avoids params name sets =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   190
      let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41228
diff changeset
   191
        val (_, ctxt') = Proof_Context.add_fixes
74523
wenzelm
parents: 74522
diff changeset
   192
          (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) lthy;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   193
        fun mk s =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   194
          let
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   195
            val t = Syntax.read_term ctxt' s;
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44045
diff changeset
   196
            val t' = fold_rev absfree params t |>
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   197
              funpow (length params) (fn Abs (_, _, t) => t)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   198
          in (t', HOLogic.dest_setT (fastype_of t)) end
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   199
          handle TERM _ =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   200
            error ("Expression " ^ quote s ^ " to be avoided in case " ^
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   201
              quote name ^ " is not a set type");
30108
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   202
        fun add_set p [] = [p]
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   203
          | add_set (t, T) ((u, U) :: ps) =
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   204
              if T = U then
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   205
                let val S = HOLogic.mk_setT T
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   206
                in (Const (\<^const_name>\<open>sup\<close>, S --> S --> S) $ u $ t, T) :: ps
30108
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   207
                end
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   208
              else (u, U) :: add_set (t, T) ps
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   209
      in
30108
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   210
        fold (mk #> add_set) sets []
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   211
      end;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   212
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   213
    val prems = map (fn (prem, name) =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   214
      let
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   215
        val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   216
        val concl = incr_boundvars 1 (Logic.strip_assums_concl prem);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   217
        val params = Logic.strip_params prem
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   218
      in
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   219
        (params,
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   220
         if null avoids then
30450
7655e6533209 HOLogic.mk_set, HOLogic.dest_set
haftmann
parents: 30364
diff changeset
   221
           map (fn (T, ts) => (HOLogic.mk_set T ts, T))
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   222
             (fold (add_binders thy 0) (prems @ [concl]) [])
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   223
         else case AList.lookup op = avoids name of
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   224
           NONE => []
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   225
         | SOME sets =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   226
             map (apfst (incr_boundvars 1)) (mk_avoids params name sets),
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   227
         prems, strip_comb (HOLogic.dest_Trueprop concl))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   228
      end) (Logic.strip_imp_prems raw_induct' ~~ induct_cases');
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   229
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   230
    val atomTs = distinct op = (maps (map snd o #2) prems);
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 78812
diff changeset
   231
    val atoms = map dest_Type_name atomTs;
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   232
    val ind_sort = if null atomTs then \<^sort>\<open>type\<close>
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 36323
diff changeset
   233
      else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn a => Sign.intern_class thy
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 36323
diff changeset
   234
        ("fs_" ^ Long_Name.base_name a)) atoms));
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42361
diff changeset
   235
    val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt');
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   236
    val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt';
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   237
    val fsT = TFree (fs_ctxt_tyname, ind_sort);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   238
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   239
    val inductive_forall_def' = Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   240
      [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   241
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   242
    fun lift_pred' t (Free (s, T)) ts =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   243
      list_comb (Free (s, fsT --> T), t :: ts);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   244
    val lift_pred = lift_pred' (Bound 0);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   245
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   246
    fun lift_prem (t as (f $ u)) =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   247
          let val (p, ts) = strip_comb t
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   248
          in
44868
92be5b32ca71 more modularization
haftmann
parents: 44689
diff changeset
   249
            if member (op =) ps p then HOLogic.mk_induct_forall fsT $
92be5b32ca71 more modularization
haftmann
parents: 44689
diff changeset
   250
              Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   251
            else lift_prem f $ lift_prem u
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   252
          end
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   253
      | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   254
      | lift_prem t = t;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   255
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   256
    fun mk_fresh (x, T) = HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   257
      (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   258
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   259
    val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   260
      let
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   261
        val params' = params @ [("y", fsT)];
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   262
        val prem = Logic.list_implies
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   263
          (map mk_fresh sets @
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   264
           map (fn prem =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   265
             if null (preds_of ps prem) then prem
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   266
             else lift_prem prem) prems,
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   267
           HOLogic.mk_Trueprop (lift_pred p ts));
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   268
      in abs_params params' prem end) prems);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   269
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   270
    val ind_vars =
80661
231d58c412b5 tuned signature: eliminate aliases;
wenzelm
parents: 80636
diff changeset
   271
      (Case_Translation.indexify_names (replicate (length atomTs) "pi") ~~
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   272
       map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   273
    val ind_Ts = rev (map snd ind_vars);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   274
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   275
    val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   276
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   277
        HOLogic.list_all (ind_vars, lift_pred p
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   278
          (map (fold_rev (NominalDatatype.mk_perm ind_Ts)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   279
            (map Bound (length atomTs downto 1))) ts)))) concls));
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   280
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   281
    val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   282
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   283
        lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls));
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   284
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   285
    val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   286
      map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32304
diff changeset
   287
          (map_filter (fn prem =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   288
             if null (preds_of ps prem) then SOME prem
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   289
             else map_term (split_conj (K o I) names) prem prem) prems, q))))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   290
        (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   291
           (NominalDatatype.fresh_star_const U T $ u $ t)) sets)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   292
             (ts ~~ binder_types (fastype_of p)) @
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   293
         map (fn (u, U) => HOLogic.mk_Trueprop (Const (\<^const_name>\<open>finite\<close>,
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   294
           HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   295
      split_list) prems |> split_list;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   296
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   297
    val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   298
    val pt2_atoms = map (fn a => Global_Theory.get_thm thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30307
diff changeset
   299
      ("pt_" ^ Long_Name.base_name a ^ "2")) atoms;
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   300
    fun eqvt_ss ctxt =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   301
      put_simpset HOL_basic_ss ctxt
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   302
        addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms)
80703
cc4ecaa8e96e tuned: prefer canonical argument order;
wenzelm
parents: 80661
diff changeset
   303
        |> Simplifier.add_proc (mk_perm_bool_simproc [\<^const_name>\<open>Fun.id\<close>])
cc4ecaa8e96e tuned: prefer canonical argument order;
wenzelm
parents: 80661
diff changeset
   304
        |> Simplifier.add_proc NominalPermeq.perm_app_simproc
cc4ecaa8e96e tuned: prefer canonical argument order;
wenzelm
parents: 80661
diff changeset
   305
        |> Simplifier.add_proc NominalPermeq.perm_fun_simproc;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   306
    val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij";
28730
71c946ce7eb9 Streamlined functions for accessing information about atoms.
berghofe
parents: 28653
diff changeset
   307
    val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms;
71c946ce7eb9 Streamlined functions for accessing information about atoms.
berghofe
parents: 28653
diff changeset
   308
    val at_insts = map (NominalAtoms.at_inst_of thy) atoms;
71c946ce7eb9 Streamlined functions for accessing information about atoms.
berghofe
parents: 28653
diff changeset
   309
    val dj_thms = maps (fn a =>
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   310
      map (NominalAtoms.dj_thm_of thy a) (remove (op =) a atoms)) atoms;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   311
    val finite_ineq = map2 (fn th => fn th' => th' RS (th RS
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   312
      @{thm pt_set_finite_ineq})) pt_insts at_insts;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   313
    val perm_set_forget =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   314
      map (fn th => th RS @{thm dj_perm_set_forget}) dj_thms;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   315
    val perm_freshs_freshs = atomTs ~~ map2 (fn th => fn th' => th' RS (th RS
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   316
      @{thm pt_freshs_freshs})) pt_insts at_insts;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   317
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   318
    fun obtain_fresh_name ts sets (T, fin) (freshs, ths1, ths2, ths3, ctxt) =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   319
      let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41228
diff changeset
   320
        val thy = Proof_Context.theory_of ctxt;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   321
        (** protect terms to avoid that fresh_star_prod_set interferes with  **)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   322
        (** pairs used in introduction rules of inductive predicate          **)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   323
        fun protect t =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   324
          let val T = fastype_of t in Const (\<^const_name>\<open>Fun.id\<close>, T --> T) $ t end;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   325
        val p = foldr1 HOLogic.mk_prod (map protect ts);
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 78812
diff changeset
   326
        val atom = dest_Type_name T;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   327
        val {at_inst, ...} = NominalAtoms.the_atom_info thy atom;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   328
        val fs_atom = Global_Theory.get_thm thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30307
diff changeset
   329
          ("fs_" ^ Long_Name.base_name atom ^ "1");
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   330
        val avoid_th = Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   331
          [SOME (Thm.global_ctyp_of thy (fastype_of p))] [SOME (Thm.global_cterm_of thy p)]
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   332
          ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding});
32202
443d5cfaba1b SUBPROOF/Obtain.result: named params;
wenzelm
parents: 32172
diff changeset
   333
        val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   334
          (fn ctxt' => EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   335
            [resolve_tac ctxt' [avoid_th] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   336
             full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   337
             full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1,
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   338
             rotate_tac 1 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   339
             REPEAT (eresolve_tac ctxt' [conjE] 1)])
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   340
          [] ctxt;
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67405
diff changeset
   341
        val (Ts1, _ :: Ts2) = chop_prefix (not o equal T) (map snd sets);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   342
        val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   343
        val (pis1, pis2) = chop (length Ts1)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   344
          (map Bound (length pTs - 1 downto 0));
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   345
        val _ $ (f $ (_ $ pi $ l) $ r) = Thm.prop_of th2
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   346
        val th2' =
54983
2c57fc1f7a8c more accurate context;
wenzelm
parents: 54895
diff changeset
   347
          Goal.prove ctxt' [] []
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 44929
diff changeset
   348
            (Logic.list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   349
               (f $ fold_rev (NominalDatatype.mk_perm (rev pTs))
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   350
                  (pis1 @ pi :: pis2) l $ r)))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   351
            (fn {context = goal_ctxt, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   352
              cut_facts_tac [th2] 1 THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   353
              full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps perm_set_forget) 1) |>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   354
          Simplifier.simplify (eqvt_ss ctxt')
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   355
      in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   356
        (freshs @ [Thm.term_of cx],
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   357
         ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt')
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   358
      end;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   359
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   360
    fun mk_ind_proof ctxt thss =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   361
      Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   362
        let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   363
          resolve_tac goal_ctxt1 [raw_induct] 1 THEN
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   364
          EVERY (maps (fn (((((_, sets, oprems, _),
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   365
              vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) =>
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   366
            [REPEAT (resolve_tac goal_ctxt1 [allI] 1), simp_tac (eqvt_ss goal_ctxt1) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   367
             SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} =>
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   368
               let
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   369
                 val (cparams', (pis, z)) =
32202
443d5cfaba1b SUBPROOF/Obtain.result: named params;
wenzelm
parents: 32172
diff changeset
   370
                   chop (length params - length atomTs - 1) (map #2 params) ||>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   371
                   (map Thm.term_of #> split_last);
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   372
                 val params' = map Thm.term_of cparams'
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   373
                 val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   374
                 val pi_sets = map (fn (t, _) =>
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   375
                   fold_rev (NominalDatatype.mk_perm []) pis t) sets';
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   376
                 val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl));
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32304
diff changeset
   377
                 val gprems1 = map_filter (fn (th, t) =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   378
                   if null (preds_of ps t) then SOME th
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   379
                   else
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   380
                     map_thm goal_ctxt2 (split_conj (K o I) names)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   381
                       (eresolve_tac goal_ctxt2 [conjunct1] 1) monos NONE th)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   382
                   (gprems ~~ oprems);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   383
                 val vc_compat_ths' = map2 (fn th => fn p =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   384
                   let
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   385
                     val th' = gprems1 MRS inst_params thy p th cparams';
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   386
                     val (h, ts) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   387
                       strip_comb (HOLogic.dest_Trueprop (Thm.concl_of th'))
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   388
                   in
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   389
                     Goal.prove goal_ctxt2 [] []
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   390
                       (HOLogic.mk_Trueprop (list_comb (h,
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   391
                          map (fold_rev (NominalDatatype.mk_perm []) pis) ts)))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   392
                       (fn {context = goal_ctxt3, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   393
                         simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   394
                          (fresh_star_bij @ finite_ineq)) 1 THEN resolve_tac goal_ctxt3 [th'] 1)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   395
                   end) vc_compat_ths vc_compat_vs;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   396
                 val (vc_compat_ths1, vc_compat_ths2) =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   397
                   chop (length vc_compat_ths - length sets) vc_compat_ths';
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   398
                 val vc_compat_ths1' = map
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   399
                   (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   400
                      (Simplifier.rewrite (eqvt_ss goal_ctxt2))))) vc_compat_ths1;
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   401
                 val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   402
                   (obtain_fresh_name ts sets)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   403
                   (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], goal_ctxt2);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   404
                 fun concat_perm pi1 pi2 =
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   405
                   let val T = fastype_of pi1
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   406
                   in if T = fastype_of pi2 then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   407
                       Const (\<^const_name>\<open>append\<close>, T --> T --> T) $ pi1 $ pi2
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   408
                     else pi2
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   409
                   end;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   410
                 val pis'' = fold_rev (concat_perm #> map) pis' pis;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   411
                 val ihyp' = inst_params thy vs_ihypt ihyp
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   412
                   (map (fold_rev (NominalDatatype.mk_perm [])
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   413
                      (pis' @ pis) #> Thm.global_cterm_of thy) params' @ [Thm.global_cterm_of thy z]);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   414
                 fun mk_pi th =
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   415
                   Simplifier.simplify
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   416
                     (put_simpset HOL_basic_ss ctxt'' addsimps [@{thm id_apply}]
80703
cc4ecaa8e96e tuned: prefer canonical argument order;
wenzelm
parents: 80661
diff changeset
   417
                       |> Simplifier.add_proc NominalDatatype.perm_simproc)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   418
                     (Simplifier.simplify (eqvt_ss ctxt'')
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   419
                       (fold_rev (mk_perm_bool ctxt'' o Thm.cterm_of ctxt'')
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   420
                         (pis' @ pis) th));
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   421
                 val gprems2 = map (fn (th, t) =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   422
                   if null (preds_of ps t) then mk_pi th
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   423
                   else
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   424
                     mk_pi (the (map_thm ctxt'' (inst_conj_all names ps (rev pis''))
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   425
                       (inst_conj_all_tac ctxt'' (length pis'')) monos (SOME t) th)))
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   426
                   (gprems ~~ oprems);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   427
                 val perm_freshs_freshs' = map (fn (th, (_, T)) =>
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   428
                   th RS the (AList.lookup op = perm_freshs_freshs T))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   429
                     (fresh_ths2 ~~ sets);
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   430
                 val th = Goal.prove ctxt'' [] []
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   431
                   (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   432
                     map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts))))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   433
                   (fn {context = goal_ctxt4, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   434
                     EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   435
                     resolve_tac goal_ctxt4 [ihyp'] 1] @
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   436
                     map (fn th => resolve_tac goal_ctxt4 [th] 1) fresh_ths3 @
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   437
                     [REPEAT_DETERM_N (length gprems)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   438
                       (simp_tac (put_simpset HOL_basic_ss goal_ctxt4
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   439
                          addsimps [inductive_forall_def']
80703
cc4ecaa8e96e tuned: prefer canonical argument order;
wenzelm
parents: 80661
diff changeset
   440
                          |> Simplifier.add_proc NominalDatatype.perm_simproc) 1 THEN
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   441
                        resolve_tac goal_ctxt4 gprems2 1)]));
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   442
                 val final = Goal.prove ctxt'' [] [] (Thm.term_of concl)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   443
                   (fn {context = goal_ctxt5, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   444
                     cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   445
                     addsimps vc_compat_ths1' @ fresh_ths1 @
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   446
                       perm_freshs_freshs') 1);
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   447
                 val final' = Proof_Context.export ctxt'' goal_ctxt2 [final];
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   448
               in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1])
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   449
                 (prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems'')))
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   450
        in
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   451
          cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   452
          REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   453
            eresolve_tac goal_ctxt [impE] 1 THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   454
            assume_tac goal_ctxt 1 THEN REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   455
            asm_full_simp_tac goal_ctxt 1)
30108
bd78f08b0ba1 Added postprocessing rules for fresh_star.
berghofe
parents: 30087
diff changeset
   456
        end) |>
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   457
        fresh_postprocess ctxt |>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74523
diff changeset
   458
        singleton (Proof_Context.export ctxt lthy);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   459
  in
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   460
    ctxt'' |>
74522
wenzelm
parents: 74282
diff changeset
   461
    Proof.theorem NONE (fn thss => fn lthy1 =>
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   462
      let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30307
diff changeset
   463
        val rec_name = space_implode "_" (map Long_Name.base_name names);
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30108
diff changeset
   464
        val rec_qualified = Binding.qualify false rec_name;
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33049
diff changeset
   465
        val ind_case_names = Rule_Cases.case_names induct_cases;
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   466
        val induct_cases' = Inductive.partition_rules' raw_induct
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   467
          (intrs ~~ induct_cases); 
74522
wenzelm
parents: 74282
diff changeset
   468
        val thss' = map (map (atomize_intr lthy1)) thss;
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   469
        val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   470
        val strong_raw_induct =
74522
wenzelm
parents: 74282
diff changeset
   471
          mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1;
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   472
        val strong_induct_atts =
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   473
          map (Attrib.internal \<^here> o K)
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   474
            [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))];
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   475
        val strong_induct =
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   476
          if length names > 1 then strong_raw_induct
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   477
          else strong_raw_induct RSN (2, rev_mp);
32304
berghofe
parents: 32303
diff changeset
   478
        val (induct_name, inducts_name) =
berghofe
parents: 32303
diff changeset
   479
          case alt_name of
32303
ba59e95a5d2b the derived induction principles can be given an explicit name
Christian Urban <urbanc@in.tum.de>
parents: 32202
diff changeset
   480
            NONE => (rec_qualified (Binding.name "strong_induct"),
ba59e95a5d2b the derived induction principles can be given an explicit name
Christian Urban <urbanc@in.tum.de>
parents: 32202
diff changeset
   481
                     rec_qualified (Binding.name "strong_inducts"))
32304
berghofe
parents: 32303
diff changeset
   482
          | SOME s => (Binding.name s, Binding.name (s ^ "s"));
74522
wenzelm
parents: 74282
diff changeset
   483
        val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   484
          ((induct_name, strong_induct_atts), [strong_induct]);
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   485
        val strong_inducts =
74522
wenzelm
parents: 74282
diff changeset
   486
          Project_Rule.projects lthy2 (1 upto length names) strong_induct'
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   487
      in
74522
wenzelm
parents: 74282
diff changeset
   488
        lthy2 |>
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   489
        Local_Theory.notes [((inducts_name, []),
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 46961
diff changeset
   490
          strong_inducts |> map (fn th => ([th],
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   491
            [Attrib.internal \<^here> (K ind_case_names),
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   492
             Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   493
      end)
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   494
      (map (map (rulify_term thy #> rpair [])) vc_compat)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   495
  end;
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   496
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   497
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   498
(* outer syntax *)
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   499
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   500
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   501
  Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>nominal_inductive2\<close>
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36692
diff changeset
   502
    "prove strong induction theorem for inductive predicate involving nominal datatypes"
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62913
diff changeset
   503
    (Parse.name -- 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   504
     Scan.option (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.name --| \<^keyword>\<open>)\<close>)) --
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   505
     (Scan.optional (\<^keyword>\<open>avoids\<close> |-- Parse.enum1 "|" (Parse.name --
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   506
      (\<^keyword>\<open>:\<close> |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) =>
32303
ba59e95a5d2b the derived induction principles can be given an explicit name
Christian Urban <urbanc@in.tum.de>
parents: 32202
diff changeset
   507
        prove_strong_ind name rule_name avoids));
28653
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   508
4593c70e228e More general, still experimental version of nominal_inductive for
berghofe
parents:
diff changeset
   509
end