src/HOL/Nominal/nominal_inductive.ML
author wenzelm
Wed, 07 Aug 2024 14:44:51 +0200
changeset 80661 231d58c412b5
parent 80636 4041e7c8059d
child 80703 cc4ecaa8e96e
permissions -rw-r--r--
tuned signature: eliminate aliases;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_inductive.ML
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     3
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
     4
Infrastructure for proving equivariance and strong induction theorems
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
     5
for inductive predicates involving nominal datatypes.
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     6
*)
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     7
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     8
signature NOMINAL_INDUCTIVE =
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
     9
sig
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
    10
  val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
    11
  val prove_eqvt: string -> string list -> local_theory -> local_theory
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    12
end
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    13
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    14
structure NominalInductive : NOMINAL_INDUCTIVE =
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    15
struct
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    16
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59936
diff changeset
    17
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
    18
val inductive_atomize = @{thms induct_atomize};
b6a1feca2ac2 use of thm-antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 33671
diff changeset
    19
val inductive_rulify = @{thms induct_rulify};
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    20
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 39557
diff changeset
    21
fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify [];
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    22
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
    23
fun atomize_conv ctxt =
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 39557
diff changeset
    24
  Raw_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
    25
    (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
    26
fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt));
24832
64cd13299d39 Conv.forall_conv: proper context;
wenzelm
parents: 24830
diff changeset
    27
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: 74522
diff changeset
    28
  (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt));
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    29
44929
haftmann
parents: 44868
diff changeset
    30
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
    31
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38795
diff changeset
    32
val fresh_prod = @{thm fresh_prod};
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    33
44689
f247fc952f31 tuned specifications
haftmann
parents: 44045
diff changeset
    34
val perm_bool = mk_meta_eq @{thm perm_bool_def};
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38795
diff changeset
    35
val perm_boolI = @{thm perm_boolI};
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    36
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    37
  (Drule.strip_imp_concl (Thm.cprop_of perm_boolI))));
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    38
60787
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
    39
fun mk_perm_bool ctxt pi th =
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
    40
  th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
    41
60359
cb8828b859a1 clarified context;
wenzelm
parents: 59940
diff changeset
    42
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
    43
  Simplifier.make_simproc \<^context>
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    44
   {name = "perm_bool",
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    45
    lhss = [\<^term>\<open>perm pi x\<close>],
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    46
    proc = fn _ => fn _ => fn ct =>
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    47
      (case Thm.term_of ct of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    48
        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
    49
          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
    50
          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
    51
      | _ => NONE),
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78806
diff changeset
    52
    identifier = []};
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    53
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    54
fun transp ([] :: _) = []
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    55
  | transp xs = map hd xs :: transp (map tl xs);
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
    56
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    57
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    58
      (Const (s, T), ts) => (case strip_type T of
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    59
        (Ts, Type (tname, _)) =>
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
    60
          (case NominalDatatype.get_nominal_datatype thy tname of
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    61
             NONE => fold (add_binders thy i) ts bs
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    62
           | SOME {descr, index, ...} => (case AList.lookup op =
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    63
                 (#3 (the (AList.lookup op = descr index))) s of
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    64
               NONE => fold (add_binders thy i) ts bs
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    65
             | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    66
                 let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs'
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    67
                 in (add_binders thy i u
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    68
                   (fold (fn (u, T) =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    69
                      if exists (fn j => j < i) (loose_bnos u) then I
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58112
diff changeset
    70
                      else insert (op aconv o apply2 fst)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    71
                        (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    72
                 end) cargs (bs, ts ~~ Ts))))
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    73
      | _ => fold (add_binders thy i) ts bs)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    74
    | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs))
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    75
  | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    76
  | add_binders thy i _ bs = bs;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
    77
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    78
fun split_conj f names (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ = (case head_of p of
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    79
      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
    80
        if member (op =) names name then SOME (f p q) else NONE
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    81
    | _ => NONE)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    82
  | split_conj _ _ _ _ = NONE;
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    83
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    84
fun strip_all [] t = t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    85
  | strip_all (_ :: xs) (Const (\<^const_name>\<open>All\<close>, _) $ Abs (s, T, t)) = strip_all xs t;
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    86
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    87
(*********************************************************************)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    88
(* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
27722
d657036e26c5 - corrected bogus comment for function inst_conj_all
berghofe
parents: 27449
diff changeset
    89
(* or    ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)            *)
d657036e26c5 - corrected bogus comment for function inst_conj_all
berghofe
parents: 27449
diff changeset
    90
(* to    R ... & id (ALL z. P z (pi_1 o ... o pi_n o t))             *)
d657036e26c5 - corrected bogus comment for function inst_conj_all
berghofe
parents: 27449
diff changeset
    91
(* or    id (ALL z. P z (pi_1 o ... o pi_n o t))                     *)
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    92
(*                                                                   *)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    93
(* where "id" protects the subformula from simplification            *)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    94
(*********************************************************************)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    95
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    96
fun inst_conj_all names ps pis (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ =
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    97
      (case head_of p of
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
    98
         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
    99
           if member (op =) names name then SOME (HOLogic.mk_conj (p,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   100
             Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   101
               (subst_bounds (pis, strip_all pis q))))
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   102
           else NONE
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   103
       | _ => NONE)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   104
  | inst_conj_all names ps pis t u =
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   105
      if member (op aconv) ps (head_of u) then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   106
        SOME (Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   107
          (subst_bounds (pis, strip_all pis t)))
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   108
      else NONE
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   109
  | inst_conj_all _ _ _ _ _ = NONE;
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   110
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   111
fun inst_conj_all_tac ctxt k = EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   112
  [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
   113
   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
   114
   simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1];
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   115
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   116
fun map_term f t u = (case f t u of
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   117
      NONE => map_term' f t u | x => x)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   118
and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   119
      (NONE, NONE) => NONE
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   120
    | (SOME t'', NONE) => SOME (t'' $ u)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   121
    | (NONE, SOME u'') => SOME (t $ u'')
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   122
    | (SOME t'', SOME u'') => SOME (t'' $ u''))
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   123
  | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   124
      NONE => NONE
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   125
    | SOME t'' => SOME (Abs (s, T, t'')))
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   126
  | map_term' _ _ _ = NONE;
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   127
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   128
(*********************************************************************)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   129
(*         Prove  F[f t]  from  F[t],  where F is monotone           *)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   130
(*********************************************************************)
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   131
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   132
fun map_thm ctxt f tac monos opt th =
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   133
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   134
    val prop = Thm.prop_of th;
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   135
    fun prove t =
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   136
      Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} =>
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   137
        EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1,
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   138
          REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)),
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   139
          REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))])
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   140
  in Option.map prove (map_term f prop (the_default prop opt)) end;
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   141
27352
8adeff7fd4bc - Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents: 27228
diff changeset
   142
val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion;
8adeff7fd4bc - Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents: 27228
diff changeset
   143
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   144
fun first_order_matchs pats objs = Thm.first_order_match
27352
8adeff7fd4bc - Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents: 27228
diff changeset
   145
  (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats),
8adeff7fd4bc - Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents: 27228
diff changeset
   146
   eta_contract_cterm (Conjunction.mk_conjunction_balanced objs));
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   147
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   148
fun first_order_mrs ths th = ths MRS
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   149
  Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   150
74522
wenzelm
parents: 74519
diff changeset
   151
fun prove_strong_ind s avoids lthy =
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   152
  let
74522
wenzelm
parents: 74519
diff changeset
   153
    val thy = Proof_Context.theory_of lthy;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   154
    val ({names, ...}, {raw_induct, intrs, elims, ...}) =
74522
wenzelm
parents: 74519
diff changeset
   155
      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
   156
    val ind_params = Inductive.params_of raw_induct;
74522
wenzelm
parents: 74519
diff changeset
   157
    val raw_induct = atomize_induct lthy raw_induct;
wenzelm
parents: 74519
diff changeset
   158
    val elims = map (atomize_induct lthy) elims;
wenzelm
parents: 74519
diff changeset
   159
    val monos = Inductive.get_monos lthy;
wenzelm
parents: 74519
diff changeset
   160
    val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy;
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   161
    val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of
22788
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   162
        [] => ()
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   163
      | xs => error ("Missing equivariance theorem for predicate(s): " ^
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   164
          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
   165
    val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the
74522
wenzelm
parents: 74519
diff changeset
   166
      (Induct.lookup_inductP lthy (hd names)))));
wenzelm
parents: 74519
diff changeset
   167
    val (raw_induct', ctxt') = lthy
70326
wenzelm
parents: 69597
diff changeset
   168
      |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct);
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   169
    val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   170
      HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   171
    val ps = map (fst o snd) concls;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   172
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58112
diff changeset
   173
    val _ = (case duplicates (op = o apply2 fst) avoids of
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   174
        [] => ()
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   175
      | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
57884
36b5691b81a5 tuned -- avoid confusion with @{assert} for system failures (exception Fail);
wenzelm
parents: 56253
diff changeset
   176
    val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse
36b5691b81a5 tuned -- avoid confusion with @{assert} for system failures (exception Fail);
wenzelm
parents: 56253
diff changeset
   177
      error ("Duplicate variable names for case " ^ quote a));
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   178
    val _ = (case subtract (op =) induct_cases (map fst avoids) of
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   179
        [] => ()
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   180
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   181
    val avoids' = if null induct_cases then replicate (length intrs) ("", [])
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   182
      else map (fn name =>
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   183
        (name, the_default [] (AList.lookup op = avoids name))) induct_cases;
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   184
    fun mk_avoids params (name, ps) =
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   185
      let val k = length params - 1
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   186
      in map (fn x => case find_index (equal x o fst) params of
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   187
          ~1 => error ("No such variable in case " ^ quote name ^
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   188
            " of inductive definition: " ^ quote x)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   189
        | i => (Bound (k - i), snd (nth params i))) ps
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   190
      end;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   191
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   192
    val prems = map (fn (prem, avoid) =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   193
      let
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   194
        val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   195
        val concl = incr_boundvars 1 (Logic.strip_assums_concl prem);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   196
        val params = Logic.strip_params prem
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   197
      in
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   198
        (params,
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   199
         fold (add_binders thy 0) (prems @ [concl]) [] @
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   200
           map (apfst (incr_boundvars 1)) (mk_avoids params avoid),
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   201
         prems, strip_comb (HOLogic.dest_Trueprop concl))
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   202
      end) (Logic.strip_imp_prems raw_induct' ~~ avoids');
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   203
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   204
    val atomTs = distinct op = (maps (map snd o #2) prems);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   205
    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
   206
      else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 78812
diff changeset
   207
        ("fs_" ^ Long_Name.base_name (dest_Type_name T))) atomTs));
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42361
diff changeset
   208
    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
   209
    val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt';
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   210
    val fsT = TFree (fs_ctxt_tyname, ind_sort);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   211
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   212
    val inductive_forall_def' = Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59586
diff changeset
   213
      [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def;
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   214
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   215
    fun lift_pred' t (Free (s, T)) ts =
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   216
      list_comb (Free (s, fsT --> T), t :: ts);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   217
    val lift_pred = lift_pred' (Bound 0);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   218
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   219
    fun lift_prem (t as (f $ u)) =
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   220
          let val (p, ts) = strip_comb t
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   221
          in
44868
92be5b32ca71 more modularization
haftmann
parents: 44689
diff changeset
   222
            if member (op =) ps p then HOLogic.mk_induct_forall fsT $
92be5b32ca71 more modularization
haftmann
parents: 44689
diff changeset
   223
              Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   224
            else lift_prem f $ lift_prem u
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   225
          end
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   226
      | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   227
      | lift_prem t = t;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   228
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   229
    fun mk_distinct [] = []
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32202
diff changeset
   230
      | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) =>
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   231
          if T = U then SOME (HOLogic.mk_Trueprop
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   232
            (HOLogic.mk_not (HOLogic.eq_const T $ x $ y)))
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   233
          else NONE) xs @ mk_distinct xs;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   234
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   235
    fun mk_fresh (x, T) = HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   236
      (NominalDatatype.fresh_const T fsT $ x $ Bound 0);
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   237
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   238
    val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   239
      let
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   240
        val params' = params @ [("y", fsT)];
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   241
        val prem = Logic.list_implies
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   242
          (map mk_fresh bvars @ mk_distinct bvars @
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   243
           map (fn prem =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   244
             if null (preds_of ps prem) then prem
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   245
             else lift_prem prem) prems,
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   246
           HOLogic.mk_Trueprop (lift_pred p ts));
29276
94b1ffec9201 qualified Term.rename_wrt_term;
wenzelm
parents: 29270
diff changeset
   247
        val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params')
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   248
      in
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 44929
diff changeset
   249
        (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem)))
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   250
      end) prems);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   251
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   252
    val ind_vars =
80661
231d58c412b5 tuned signature: eliminate aliases;
wenzelm
parents: 80636
diff changeset
   253
      (Case_Translation.indexify_names (replicate (length atomTs) "pi") ~~
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   254
       map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   255
    val ind_Ts = rev (map snd ind_vars);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   256
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   257
    val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   258
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   259
        HOLogic.list_all (ind_vars, lift_pred p
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   260
          (map (fold_rev (NominalDatatype.mk_perm ind_Ts)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   261
            (map Bound (length atomTs downto 1))) ts)))) concls));
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   262
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   263
    val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   264
      (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   265
        lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls));
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   266
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   267
    val vc_compat = map (fn (params, bvars, prems, (p, ts)) =>
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 44929
diff changeset
   268
      map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32202
diff changeset
   269
          (map_filter (fn prem =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   270
             if null (preds_of ps prem) then SOME prem
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   271
             else map_term (split_conj (K o I) names) prem prem) prems, q))))
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   272
        (mk_distinct bvars @
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   273
         maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   274
           (NominalDatatype.fresh_const U T $ u $ t)) bvars)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   275
             (ts ~~ binder_types (fastype_of p)))) prems;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   276
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
   277
    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
   278
    val pt2_atoms = map (fn aT => Global_Theory.get_thm thy
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 78812
diff changeset
   279
      ("pt_" ^ Long_Name.base_name (dest_Type_name aT) ^ "2")) atomTs;
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   280
    fun eqvt_ss ctxt =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   281
      put_simpset HOL_basic_ss ctxt
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   282
        addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   283
        addsimprocs [mk_perm_bool_simproc [\<^const_name>\<open>Fun.id\<close>],
78806
aca84704d46f more standard simproc_setup using ML antiquotation;
wenzelm
parents: 78097
diff changeset
   284
          NominalPermeq.perm_app_simproc, 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
   285
    val fresh_bij = Global_Theory.get_thms thy "fresh_bij";
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
   286
    val perm_bij = Global_Theory.get_thms thy "perm_bij";
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
   287
    val fs_atoms = map (fn aT => Global_Theory.get_thm thy
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 78812
diff changeset
   288
      ("fs_" ^ Long_Name.base_name (dest_Type_name aT) ^ "1")) atomTs;
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
   289
    val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'";
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
   290
    val fresh_atm = Global_Theory.get_thms thy "fresh_atm";
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
   291
    val swap_simps = Global_Theory.get_thms thy "swap_simps";
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
   292
    val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh";
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   293
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   294
    fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) =
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   295
      let
27722
d657036e26c5 - corrected bogus comment for function inst_conj_all
berghofe
parents: 27449
diff changeset
   296
        (** protect terms to avoid that fresh_prod interferes with  **)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   297
        (** pairs used in introduction rules of inductive predicate **)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   298
        fun protect t =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   299
          let val T = fastype_of t in Const (\<^const_name>\<open>Fun.id\<close>, T --> T) $ t end;
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   300
        val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1);
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   301
        val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   302
            (HOLogic.exists_const T $ Abs ("x", T,
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   303
              NominalDatatype.fresh_const T (fastype_of p) $
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   304
                Bound 0 $ p)))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   305
          (fn {context = goal_ctxt, ...} => EVERY
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   306
            [resolve_tac goal_ctxt exists_fresh' 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   307
             resolve_tac goal_ctxt fs_atoms 1]);
32202
443d5cfaba1b SUBPROOF/Obtain.result: named params;
wenzelm
parents: 32172
diff changeset
   308
        val (([(_, cx)], ths), ctxt') = Obtain.result
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   309
          (fn goal_ctxt => EVERY
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   310
            [eresolve_tac goal_ctxt [exE] 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   311
             full_simp_tac (put_simpset HOL_ss goal_ctxt addsimps (fresh_prod :: fresh_atm)) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   312
             full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps [@{thm id_apply}]) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   313
             REPEAT (eresolve_tac goal_ctxt [conjE] 1)])
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   314
          [ex] ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   315
      in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   316
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   317
    fun mk_ind_proof ctxt thss =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   318
      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: 74522
diff changeset
   319
        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: 74522
diff changeset
   320
          resolve_tac goal_ctxt1 [raw_induct] 1 THEN
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   321
          EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) =>
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   322
            [REPEAT (resolve_tac goal_ctxt1 [allI] 1),
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   323
             simp_tac (eqvt_ss goal_ctxt1) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   324
             SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} =>
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   325
               let
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   326
                 val (params', (pis, z)) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   327
                   chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||>
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   328
                   split_last;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   329
                 val bvars' = map
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   330
                   (fn (Bound i, T) => (nth params' (length params' - i), T)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   331
                     | (t, T) => (t, T)) bvars;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   332
                 val pi_bvars = map (fn (t, _) =>
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   333
                   fold_rev (NominalDatatype.mk_perm []) pis t) bvars';
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   334
                 val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl));
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   335
                 val (freshs1, freshs2, ctxt') = fold
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   336
                   (obtain_fresh_name (ts @ pi_bvars))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   337
                   (map snd bvars') ([], [], goal_ctxt2);
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   338
                 val freshs2' = NominalDatatype.mk_not_sym freshs2;
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   339
                 val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1);
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   340
                 fun concat_perm pi1 pi2 =
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   341
                   let val T = fastype_of pi1
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   342
                   in if T = fastype_of pi2 then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   343
                       Const (\<^const_name>\<open>append\<close>, T --> T --> T) $ pi1 $ pi2
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   344
                     else pi2
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   345
                   end;
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   346
                 val pis'' = fold (concat_perm #> map) pis' pis;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   347
                 val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   348
                   (Vartab.empty, Vartab.empty);
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   349
                 val ihyp' = Thm.instantiate (TVars.empty,
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   350
                   Vars.make (map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t))
32035
8e77b6a250d5 tuned/modernized Envir.subst_XXX;
wenzelm
parents: 31938
diff changeset
   351
                   (map (Envir.subst_term env) vs ~~
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   352
                    map (fold_rev (NominalDatatype.mk_perm [])
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70326
diff changeset
   353
                      (rev pis' @ pis)) params' @ [z]))) ihyp;
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   354
                 fun mk_pi th =
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   355
                   Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   356
                       addsimprocs [NominalDatatype.perm_simproc])
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   357
                     (Simplifier.simplify (eqvt_ss ctxt')
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   358
                       (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt')
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   359
                         (rev pis' @ pis) th));
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   360
                 val (gprems1, gprems2) = split_list
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   361
                   (map (fn (th, t) =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   362
                      if null (preds_of ps t) then (SOME th, mk_pi th)
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   363
                      else
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   364
                        (map_thm ctxt' (split_conj (K o I) names)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   365
                           (eresolve_tac ctxt' [conjunct1] 1) monos NONE th,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   366
                         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: 74522
diff changeset
   367
                           (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th))))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32202
diff changeset
   368
                      (gprems ~~ oprems)) |>> map_filter I;
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   369
                 val vc_compat_ths' = map (fn th =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   370
                   let
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   371
                     val th' = first_order_mrs gprems1 th;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   372
                     val (bop, lhs, rhs) = (case Thm.concl_of th' of
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   373
                         _ $ (fresh $ lhs $ rhs) =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   374
                           (fn t => fn u => fresh $ t $ u, lhs, rhs)
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   375
                       | _ $ (_ $ (_ $ lhs $ rhs)) =>
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   376
                           (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs));
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   377
                     val th'' = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   378
                         (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs)
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   379
                            (fold_rev (NominalDatatype.mk_perm []) pis rhs)))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   380
                       (fn {context = goal_ctxt3, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   381
                         simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   382
                          (fresh_bij @ perm_bij)) 1 THEN resolve_tac goal_ctxt3 [th'] 1)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   383
                   in Simplifier.simplify (eqvt_ss ctxt' addsimps fresh_atm) th'' end)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   384
                     vc_compat_ths;
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   385
                 val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths';
27449
4880da911af0 Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents: 27353
diff changeset
   386
                 (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **)
4880da911af0 Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents: 27353
diff changeset
   387
                 (** we have to pre-simplify the rewrite rules                   **)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   388
                 val swap_simps_simpset = put_simpset HOL_ss ctxt' addsimps swap_simps @
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   389
                    map (Simplifier.simplify (put_simpset HOL_ss ctxt' addsimps swap_simps))
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   390
                      (vc_compat_ths'' @ freshs2');
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   391
                 val th = Goal.prove ctxt' [] []
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   392
                   (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   393
                     map (fold (NominalDatatype.mk_perm []) pis') (tl ts))))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   394
                   (fn {context = goal_ctxt4, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   395
                     EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1,
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   396
                     resolve_tac goal_ctxt4 [ihyp'] 1,
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   397
                     REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   398
                       (simp_tac swap_simps_simpset 1),
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   399
                     REPEAT_DETERM_N (length gprems)
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   400
                       (simp_tac (put_simpset HOL_basic_ss goal_ctxt4
27847
0dffedf9aff5 Changed proof of strong induction rule to avoid infinite loop
berghofe
parents: 27722
diff changeset
   401
                          addsimps [inductive_forall_def']
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   402
                          addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   403
                        resolve_tac goal_ctxt4 gprems2 1)]));
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   404
                 val final = Goal.prove ctxt' [] [] (Thm.term_of concl)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   405
                   (fn {context = goal_ctxt5, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   406
                     cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   407
                     addsimps vc_compat_ths'' @ freshs2' @
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   408
                       perm_fresh_fresh @ fresh_atm) 1);
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   409
                 val final' = Proof_Context.export ctxt' goal_ctxt2 [final];
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   410
               in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1])
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   411
                 (prems ~~ thss ~~ ihyps ~~ prems'')))
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   412
        in
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   413
          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: 74522
diff changeset
   414
          REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   415
            eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   416
            REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   417
            asm_full_simp_tac goal_ctxt 1)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   418
        end) |> singleton (Proof_Context.export ctxt lthy);
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   419
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   420
    (** strong case analysis rule **)
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   421
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   422
    val cases_prems = map (fn ((name, avoids), rule) =>
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   423
      let
74522
wenzelm
parents: 74519
diff changeset
   424
        val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] lthy;
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   425
        val prem :: prems = Logic.strip_imp_prems rule';
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   426
        val concl = Logic.strip_imp_concl rule'
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   427
      in
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   428
        (prem,
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   429
         List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params),
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   430
         concl,
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   431
         fold_map (fn (prem, (_, avoid)) => fn ctxt =>
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   432
           let
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   433
             val prems = Logic.strip_assums_hyp prem;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   434
             val params = Logic.strip_params prem;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   435
             val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid;
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   436
             fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) =
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   437
               if member (op = o apsnd fst) bnds (Bound i) then
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   438
                 let
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   439
                   val ([s'], ctxt') = Variable.variant_fixes [s] ctxt;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   440
                   val t = Free (s', T)
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   441
                 in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   442
               else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts);
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   443
             val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   444
               (0, 0, ctxt, [], [], [], [])
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   445
           in
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   446
             ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt')
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   447
           end) (prems ~~ avoids) ctxt')
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   448
      end)
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   449
        (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   450
         elims);
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   451
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   452
    val cases_prems' =
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   453
      map (fn (prem, args, concl, (prems, _)) =>
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   454
        let
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   455
          fun mk_prem (ps, [], _, prems) =
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 44929
diff changeset
   456
                Logic.list_all (ps, Logic.list_implies (prems, concl))
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   457
            | mk_prem (ps, qs, _, prems) =
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 44929
diff changeset
   458
                Logic.list_all (ps, Logic.mk_implies
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   459
                  (Logic.list_implies
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   460
                    (mk_distinct qs @
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   461
                     maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   462
                      (NominalDatatype.fresh_const T (fastype_of u) $ t $ u))
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   463
                        args) qs,
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   464
                     HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   465
                       (map HOLogic.dest_Trueprop prems))),
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   466
                   concl))
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   467
          in map mk_prem prems end) cases_prems;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   468
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   469
    fun cases_eqvt_simpset ctxt =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   470
      put_simpset HOL_ss ctxt
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   471
        addsimps eqvt_thms @ swap_simps @ perm_pi_simp
78806
aca84704d46f more standard simproc_setup using ML antiquotation;
wenzelm
parents: 78097
diff changeset
   472
        addsimprocs [NominalPermeq.perm_app_simproc, NominalPermeq.perm_fun_simproc];
27352
8adeff7fd4bc - Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents: 27228
diff changeset
   473
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   474
    fun simp_fresh_atm ctxt =
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   475
      Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps fresh_atm);
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   476
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   477
    fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt))),
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   478
        prems') =
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   479
      (name, Goal.prove ctxt [] (prem :: prems') concl
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26568
diff changeset
   480
        (fn {prems = hyp :: hyps, context = ctxt1} =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   481
        EVERY (resolve_tac ctxt1 [hyp RS elim] 1 ::
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   482
          map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) =>
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   483
            SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} =>
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   484
              if null qs then
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   485
                resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   486
              else
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   487
                let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   488
                  val params' = map (Thm.term_of o #2 o nth (rev params)) is;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   489
                  val tab = params' ~~ map fst qs;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   490
                  val (hyps1, hyps2) = chop (length args) case_hyps;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   491
                  (* turns a = t and [x1 # t, ..., xn # t] *)
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   492
                  (* into [x1 # a, ..., xn # a]            *)
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   493
                  fun inst_fresh th' ths =
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   494
                    let val (ths1, ths2) = chop (length qs) ths
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   495
                    in
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   496
                      (map (fn th =>
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   497
                         let
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   498
                           val (cf, ct) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   499
                             Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th));
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   500
                           val arg_cong' = Thm.instantiate'
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   501
                             [SOME (Thm.ctyp_of_cterm ct)]
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   502
                             [NONE, SOME ct, SOME cf] (arg_cong RS iffD2);
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   503
                           val inst = Thm.first_order_match (ct,
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   504
                             Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th')))
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   505
                         in [th', th] MRS Thm.instantiate inst arg_cong'
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   506
                         end) ths1,
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   507
                       ths2)
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   508
                    end;
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   509
                  val (vc_compat_ths1, vc_compat_ths2) =
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   510
                    chop (length vc_compat_ths - length args * length qs)
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   511
                      (map (first_order_mrs hyps2) vc_compat_ths);
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   512
                  val vc_compat_ths' =
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   513
                    NominalDatatype.mk_not_sym vc_compat_ths1 @
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   514
                    flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2));
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   515
                  val (freshs1, freshs2, ctxt3) = fold
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   516
                    (obtain_fresh_name (args @ map fst qs @ params'))
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   517
                    (map snd qs) ([], [], ctxt2);
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   518
                  val freshs2' = NominalDatatype.mk_not_sym freshs2;
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   519
                  val pis = map (NominalDatatype.perm_of_pair)
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   520
                    ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1));
60787
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   521
                  val mk_pis = fold_rev (mk_perm_bool ctxt3) (map (Thm.cterm_of ctxt3) pis);
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59586
diff changeset
   522
                  val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   523
                     (fn x as Free _ =>
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
   524
                           if member (op =) args x then x
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   525
                           else (case AList.lookup op = tab x of
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   526
                             SOME y => y
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   527
                           | NONE => fold_rev (NominalDatatype.mk_perm []) pis x)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   528
                       | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps));
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   529
                  val inst = Thm.first_order_match (Thm.dest_arg
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   530
                    (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   531
                  val th = Goal.prove ctxt3 [] [] (Thm.term_of concl)
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   532
                    (fn {context = ctxt4, ...} =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   533
                       resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   534
                       SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} =>
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   535
                         let
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   536
                           val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps;
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   537
                           val case_simpset = cases_eqvt_simpset ctxt5 addsimps freshs2' @
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   538
                             map (simp_fresh_atm ctxt5) (vc_compat_ths' @ fresh_hyps');
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   539
                           val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   540
                           val hyps1' = map
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   541
                             (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1;
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   542
                           val hyps2' = map
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   543
                             (mk_pis #> Simplifier.simplify case_simpset) hyps2;
27449
4880da911af0 Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents: 27353
diff changeset
   544
                           val case_hyps' = hyps1' @ hyps2'
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   545
                         in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50771
diff changeset
   546
                           simp_tac case_simpset 1 THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   547
                           REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   548
                             resolve_tac ctxt5 case_hyps' 1)
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   549
                         end) ctxt4 1)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41228
diff changeset
   550
                  val final = Proof_Context.export ctxt3 ctxt2 [th]
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   551
                in resolve_tac ctxt2 final 1 end) ctxt1 1)
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   552
                  (thss ~~ hyps ~~ prems))) |>
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   553
                  singleton (Proof_Context.export ctxt lthy))
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   554
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   555
  in
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   556
    ctxt'' |>
74522
wenzelm
parents: 74519
diff changeset
   557
    Proof.theorem NONE (fn thss => fn lthy1 =>
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   558
      let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   559
        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: 30087
diff changeset
   560
        val rec_qualified = Binding.qualify false rec_name;
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33049
diff changeset
   561
        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
   562
        val induct_cases' = Inductive.partition_rules' raw_induct
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   563
          (intrs ~~ induct_cases); 
74522
wenzelm
parents: 74519
diff changeset
   564
        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
   565
        val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   566
        val strong_raw_induct =
74522
wenzelm
parents: 74519
diff changeset
   567
          mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1;
wenzelm
parents: 74519
diff changeset
   568
        val strong_cases = map (mk_cases_proof ##> Inductive.rulify lthy1)
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   569
          (thsss ~~ elims ~~ cases_prems ~~ cases_prems');
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
   570
        val strong_induct_atts =
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   571
          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
   572
            [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))];
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   573
        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
   574
          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
   575
          else strong_raw_induct RSN (2, rev_mp);
74522
wenzelm
parents: 74519
diff changeset
   576
        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
   577
          ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]);
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   578
        val strong_inducts =
74522
wenzelm
parents: 74519
diff changeset
   579
          Project_Rule.projects lthy1 (1 upto length names) strong_induct';
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   580
      in
74522
wenzelm
parents: 74519
diff changeset
   581
        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
   582
        Local_Theory.notes
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
   583
          [((rec_qualified (Binding.name "strong_inducts"), []),
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
   584
            strong_inducts |> map (fn th => ([th],
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   585
              [Attrib.internal \<^here> (K ind_case_names),
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   586
               Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   587
        Local_Theory.notes (map (fn ((name, elim), (_, cases)) =>
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   588
            ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"),
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   589
              [Attrib.internal \<^here> (K (Rule_Cases.case_names (map snd cases))),
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 74524
diff changeset
   590
               Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])]))
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   591
          (strong_cases ~~ induct_cases')) |> snd
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   592
      end)
24570
621b60b1df00 Generalized equivariance and nominal_inductive commands to
berghofe
parents: 23531
diff changeset
   593
      (map (map (rulify_term thy #> rpair [])) vc_compat)
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   594
  end;
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   595
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   596
fun prove_eqvt s xatoms lthy =
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   597
  let
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   598
    val thy = Proof_Context.theory_of lthy;
22788
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   599
    val ({names, ...}, {raw_induct, intrs, elims, ...}) =
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   600
      Inductive.the_inductive_global lthy (Sign.intern_const thy s);
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   601
    val raw_induct = atomize_induct lthy raw_induct;
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   602
    val elims = map (atomize_induct lthy) elims;
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   603
    val intrs = map (atomize_intr lthy) intrs;
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   604
    val monos = Inductive.get_monos lthy;
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   605
    val intrs' = Inductive.unpartition_rules intrs
22788
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   606
      (map (fn (((s, ths), (_, k)), th) =>
60362
befdc10ebb42 clarified context;
wenzelm
parents: 60359
diff changeset
   607
           (s, ths ~~ Inductive.infer_intro_vars thy th k ths))
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   608
         (Inductive.partition_rules raw_induct intrs ~~
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   609
          Inductive.arities_of raw_induct ~~ elims));
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
   610
    val k = length (Inductive.params_of raw_induct);
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   611
    val atoms' = NominalAtoms.atoms_of thy;
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   612
    val atoms =
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   613
      if null xatoms then atoms' else
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   614
      let val atoms = map (Sign.intern_type thy) xatoms
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   615
      in
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   616
        (case duplicates op = atoms of
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   617
             [] => ()
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   618
           | xs => error ("Duplicate atoms: " ^ commas xs);
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   619
         case subtract (op =) atoms' atoms of
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   620
             [] => ()
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   621
           | xs => error ("No such atoms: " ^ commas xs);
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   622
         atoms)
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22544
diff changeset
   623
      end;
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
   624
    val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp";
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   625
    val (([t], [pi]), ctxt1) = lthy |>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   626
      Variable.import_terms false [Thm.concl_of raw_induct] ||>>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   627
      Variable.variant_fixes ["pi"];
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   628
    fun eqvt_simpset ctxt = put_simpset HOL_basic_ss ctxt addsimps
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   629
      (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs
54991
1169c65e9698 clarified context;
wenzelm
parents: 54983
diff changeset
   630
      [mk_perm_bool_simproc names,
78806
aca84704d46f more standard simproc_setup using ML antiquotation;
wenzelm
parents: 78097
diff changeset
   631
       NominalPermeq.perm_app_simproc, NominalPermeq.perm_fun_simproc];
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   632
    val ps = map (fst o HOLogic.dest_imp)
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   633
      (HOLogic.dest_conj (HOLogic.dest_Trueprop t));
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   634
    fun eqvt_tac ctxt pi (intr, vs) st =
22544
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   635
      let
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   636
        fun eqvt_err s =
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   637
          let val ([t], ctxt') = Variable.import_terms true [Thm.prop_of intr] ctxt
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   638
          in error ("Could not prove equivariance for introduction rule\n" ^
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   639
            Syntax.string_of_term ctxt' t ^ "\n" ^ s)
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   640
          end;
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   641
        val res = SUBPROOF (fn {context = goal_ctxt, prems, params, ...} =>
22788
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   642
          let
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   643
            val prems' = map (fn th => the_default th (map_thm goal_ctxt
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   644
              (split_conj (K I) names) (eresolve_tac goal_ctxt [conjunct2] 1) monos NONE th)) prems;
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   645
            val prems'' = map (fn th => Simplifier.simplify (eqvt_simpset goal_ctxt)
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   646
              (mk_perm_bool goal_ctxt (Thm.cterm_of goal_ctxt pi) th)) prems';
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   647
            val intr' = infer_instantiate goal_ctxt (map (#1 o dest_Var) vs ~~
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   648
               map (Thm.cterm_of goal_ctxt o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params)
22788
3038bd211582 eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents: 22755
diff changeset
   649
               intr
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   650
          in (resolve_tac goal_ctxt [intr'] THEN_ALL_NEW (TRY o resolve_tac goal_ctxt prems'')) 1
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   651
          end) ctxt 1 st
22544
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   652
      in
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   653
        case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   654
          NONE => eqvt_err ("Rule does not match goal\n" ^
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   655
            Syntax.string_of_term ctxt (hd (Thm.prems_of st)))
22544
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   656
        | SOME (th, _) => Seq.single th
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   657
      end;
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   658
    val thss = map (fn atom =>
25824
f56dd9745d1b Implemented proof of strong case analysis rule.
berghofe
parents: 24867
diff changeset
   659
      let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, [])))
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   660
      in map (fn th => zero_var_indexes (th RS mp))
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   661
        (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt1 [] []
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   662
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   663
            let
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   664
              val (h, ts) = strip_comb p;
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   665
              val (ts1, ts2) = chop k ts
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   666
            in
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   667
              HOLogic.mk_imp (p, list_comb (h, ts1 @
31938
f193d95b4632 nominal.ML is nominal_datatype.ML
haftmann
parents: 31723
diff changeset
   668
                map (NominalDatatype.mk_perm [] pi') ts2))
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   669
            end) ps)))
74524
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   670
          (fn {context = goal_ctxt, ...} =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   671
            EVERY (resolve_tac goal_ctxt [raw_induct] 1 :: map (fn intr_vs =>
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   672
              full_simp_tac (eqvt_simpset goal_ctxt) 1 THEN
8ee3d5d3c1bf more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
wenzelm
parents: 74522
diff changeset
   673
              eqvt_tac goal_ctxt pi' intr_vs) intrs')) |>
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   674
          singleton (Proof_Context.export ctxt1 lthy)))
22544
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   675
      end) atoms
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   676
  in
74519
fc65e39ca170 clarified context;
wenzelm
parents: 74282
diff changeset
   677
    lthy |>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   678
    Local_Theory.notes (map (fn (name, ths) =>
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   679
        ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"),
78097
2ea20bb1493c tuned: more antiquotations;
wenzelm
parents: 78095
diff changeset
   680
          @{attributes [eqvt]}), [(ths, [])]))
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   681
      (names ~~ transp thss)) |> snd
22544
549615dcd4f2 - Improved error messages in equivariance proof
berghofe
parents: 22530
diff changeset
   682
  end;
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   683
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   684
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   685
(* outer syntax *)
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   686
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24861
diff changeset
   687
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   688
  Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>nominal_inductive\<close>
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36692
diff changeset
   689
    "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   690
    (Parse.name -- Scan.optional (\<^keyword>\<open>avoids\<close> |-- Parse.and_list1 (Parse.name --
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   691
      (\<^keyword>\<open>:\<close> |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) =>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   692
        prove_strong_ind name avoids));
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   693
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24861
diff changeset
   694
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   695
  Outer_Syntax.local_theory \<^command_keyword>\<open>equivariance\<close>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   696
    "prove equivariance for inductive predicate involving nominal datatypes"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
   697
    (Parse.name -- Scan.optional (\<^keyword>\<open>[\<close> |-- Parse.list1 Parse.name --| \<^keyword>\<open>]\<close>) [] >>
30087
a780642a9c9c nominal_inductive and equivariance now work on local_theory.
berghofe
parents: 29585
diff changeset
   698
      (fn (name, atoms) => prove_eqvt name atoms));
22530
c192c5d1a6f2 Implemented proof of strong induction rule.
berghofe
parents: 22313
diff changeset
   699
22313
1a507b463f50 First steps towards strengthening of induction rules for
berghofe
parents:
diff changeset
   700
end