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