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