src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Fri, 14 Sep 2012 22:23:11 +0200
changeset 49379 7860bd1429f4
parent 49378 19237e465055
child 49380 f4b0121b13ab
permissions -rw-r--r--
provide more guidance, exploiting our knowledge of the goal
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     3
    Copyright   2012
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     4
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     5
Tactics for the LFP/GFP sugar.
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     6
*)
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     7
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     8
signature BNF_FP_SUGAR_TACTICS =
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     9
sig
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    10
  val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
49226
510c6d4a73ec fixed and enabled iterator/recursor theorems
blanchet
parents: 49218
diff changeset
    11
  val mk_coiter_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    12
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
    13
  val mk_fld_iff_unf_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
    14
    tactic
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    15
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
49376
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
    16
  val mk_induct_tac: Proof.context -> int list -> int list list ->
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
    17
    ((int * int) * (int * int)) list list list -> thm list -> thm -> thm list -> thm list list ->
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
    18
    tactic
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    19
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
49226
510c6d4a73ec fixed and enabled iterator/recursor theorems
blanchet
parents: 49218
diff changeset
    20
  val mk_iter_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    21
end;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    22
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    23
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    24
struct
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    25
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    26
open BNF_Tactics
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    27
open BNF_Util
49264
9059e0dbdbc1 implemented and use "mk_sum_casesN_balanced"
blanchet
parents: 49263
diff changeset
    28
open BNF_FP_Util
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    29
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    30
val meta_mp = @{thm meta_mp};
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    31
val meta_spec = @{thm meta_spec};
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    32
49370
be6e749fd003 fixed variable exporting problem
blanchet
parents: 49368
diff changeset
    33
fun smash_spurious_fs lthy thm =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    34
  let
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    35
    val spurious_fs =
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    36
      Term.add_vars (prop_of thm) []
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    37
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    38
    val cxs =
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    39
      map (fn s as (_, T) =>
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    40
        (certify lthy (Var s), certify lthy (id_abs (domain_type T)))) spurious_fs;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    41
  in
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    42
    Drule.cterm_instantiate cxs thm
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    43
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    44
49370
be6e749fd003 fixed variable exporting problem
blanchet
parents: 49368
diff changeset
    45
val smash_spurious_fs_tac = PRIMITIVE o smash_spurious_fs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    46
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    47
fun mk_case_tac ctxt n k m case_def ctr_def unf_fld =
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    48
  Local_Defs.unfold_tac ctxt [case_def, ctr_def, unf_fld] THEN
49264
9059e0dbdbc1 implemented and use "mk_sum_casesN_balanced"
blanchet
parents: 49263
diff changeset
    49
  (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    50
   REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    51
   rtac refl) 1;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    52
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    53
fun mk_exhaust_tac ctxt n ctr_defs fld_iff_unf sumEN' =
49263
669a820ef213 fixed general case of "mk_sumEN_balanced"
blanchet
parents: 49262
diff changeset
    54
  Local_Defs.unfold_tac ctxt (fld_iff_unf :: ctr_defs) THEN rtac sumEN' 1 THEN
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    55
  Local_Defs.unfold_tac ctxt @{thms all_prod_eq} THEN
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    56
  EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    57
    etac meta_mp, atac]) (1 upto n)) 1;
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    58
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    59
fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    60
  (rtac iffI THEN'
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    61
   EVERY' (map3 (fn cTs => fn cx => fn th =>
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    62
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    63
     SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    64
     atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    65
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    66
fun mk_half_distinct_tac ctxt fld_inject ctr_defs =
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    67
  Local_Defs.unfold_tac ctxt (fld_inject :: @{thms sum.inject} @ ctr_defs) THEN
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    68
  rtac @{thm sum.distinct(1)} 1;
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    69
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    70
fun mk_inject_tac ctxt ctr_def fld_inject =
49135
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49130
diff changeset
    71
  Local_Defs.unfold_tac ctxt [ctr_def] THEN rtac (fld_inject RS ssubst) 1 THEN
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49130
diff changeset
    72
  Local_Defs.unfold_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    73
49211
239a4fa29ddf define corecursors
blanchet
parents: 49207
diff changeset
    74
val iter_like_thms =
49233
7f412734fbb3 fixed and reenabled "corecs" theorems
blanchet
parents: 49232
diff changeset
    75
  @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
7f412734fbb3 fixed and reenabled "corecs" theorems
blanchet
parents: 49232
diff changeset
    76
      split_conv};
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
    77
49229
d5717b5e2217 use map_id, not map_id', to allow better composition
blanchet
parents: 49226
diff changeset
    78
fun mk_iter_like_tac pre_map_defs map_ids iter_like_defs fld_iter_like ctr_def ctxt =
d5717b5e2217 use map_id, not map_id', to allow better composition
blanchet
parents: 49226
diff changeset
    79
  Local_Defs.unfold_tac ctxt (ctr_def :: fld_iter_like :: iter_like_defs @ pre_map_defs @ map_ids @
49232
9ea11f0c53e4 fixed and enabled generation of "coiters" theorems, including the recursive case
blanchet
parents: 49229
diff changeset
    80
    iter_like_thms) THEN Local_Defs.unfold_tac ctxt @{thms id_def} THEN rtac refl 1;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
    81
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
    82
val coiter_like_ss = ss_only @{thms if_True if_False};
49276
59fa53ed7507 finished splitting sum types for corecursors
blanchet
parents: 49274
diff changeset
    83
val coiter_like_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
    84
49229
d5717b5e2217 use map_id, not map_id', to allow better composition
blanchet
parents: 49226
diff changeset
    85
fun mk_coiter_like_tac coiter_like_defs map_ids fld_unf_coiter_like pre_map_def ctr_def ctxt =
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
    86
  Local_Defs.unfold_tac ctxt (ctr_def :: coiter_like_defs) THEN
49226
510c6d4a73ec fixed and enabled iterator/recursor theorems
blanchet
parents: 49218
diff changeset
    87
  subst_tac ctxt [fld_unf_coiter_like] 1 THEN asm_simp_tac coiter_like_ss 1 THEN
49232
9ea11f0c53e4 fixed and enabled generation of "coiters" theorems, including the recursive case
blanchet
parents: 49229
diff changeset
    88
  Local_Defs.unfold_tac ctxt (pre_map_def :: coiter_like_thms @ map_ids) THEN
49235
f9fc2b64c599 fixed bug with one-value codatatype "codata 'a dead_foo = A"
blanchet
parents: 49233
diff changeset
    89
  Local_Defs.unfold_tac ctxt @{thms id_def} THEN
49362
1271aca16aed make tactic more robust in the case where "asm_simp_tac" already finishes the job
blanchet
parents: 49276
diff changeset
    90
  TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
    91
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    92
fun mk_induct_prelude_tac ctxt ctr_defs fld_induct' =
49370
be6e749fd003 fixed variable exporting problem
blanchet
parents: 49368
diff changeset
    93
  Local_Defs.unfold_tac ctxt ctr_defs THEN rtac fld_induct' 1 THEN smash_spurious_fs_tac ctxt;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    94
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    95
fun mk_induct_prepare_prem_tac n m k =
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    96
  EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    97
    REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    98
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
    99
fun mk_induct_prepare_prem_prems_tac 0 = all_tac
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   100
  | mk_induct_prepare_prem_prems_tac r =
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   101
    REPEAT_DETERM_N r ((rotate_tac ~1) 1 THEN dtac meta_mp 1 THEN
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   102
      defer_tac 2 THEN PRIMITIVE (Thm.permute_prems 0 ~1) THEN rotate_tac 1 1) THEN
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   103
    PRIMITIVE Raw_Simplifier.norm_hhf;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   104
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   105
val induct_prem_prem_thms =
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   106
  @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex UN_insert Un_assoc[symmetric] Un_empty_left
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   107
      Un_empty_right Union_Un_distrib collect_def[abs_def] fst_conv image_def o_apply snd_conv
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   108
      snd_prod_fun sum.cases sup_bot_right fst_map_pair map_pair_simp sum_map.simps};
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   109
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   110
(* These rules interfere with the "set_natural'" properties of "sum" and "prod", so we explicitly
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   111
   delay them. *)
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   112
val induct_prem_prem_thms_delayed =
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   113
  @{thms fsts_def[abs_def] snds_def[abs_def] sum_setl_def[abs_def] sum_setr_def[abs_def]};
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   114
49379
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   115
fun mk_induct_prem_prem_endgame_tac _ 0 = atac ORELSE' rtac @{thm singletonI}
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   116
  | mk_induct_prem_prem_endgame_tac ctxt qq =
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   117
    REPEAT_DETERM_N qq o
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   118
      (SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms Union_iff bex_simps(6)}) THEN'
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   119
       etac @{thm induct_set_step}) THEN'
7860bd1429f4 provide more guidance, exploiting our knowledge of the goal
blanchet
parents: 49378
diff changeset
   120
    (atac ORELSE' blast_tac ctxt);
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   121
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   122
(* TODO: Get rid of the "blast_tac" *)
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   123
fun mk_induct_discharge_prem_prems_tac ctxt nn ixs set_natural's pre_set_defs =
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   124
  EVERY' (maps (fn ((pp, jj), (qq, kk)) =>
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   125
    [select_prem_tac nn (dtac meta_spec) (nn - kk + 1), rotate_tac ~1,(*###*) etac meta_mp,
49375
993677c1cf2a tuned code before fixing "mk_induct_discharge_prem_prems_tac"
blanchet
parents: 49372
diff changeset
   126
     SELECT_GOAL (Local_Defs.unfold_tac ctxt pre_set_defs), (* ### why on a line of its own? *)
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   127
     SELECT_GOAL (Local_Defs.unfold_tac ctxt (set_natural's @ induct_prem_prem_thms)),
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   128
     SELECT_GOAL (Local_Defs.unfold_tac ctxt
49372
c62165188971 polished the induction
blanchet
parents: 49370
diff changeset
   129
       (induct_prem_prem_thms_delayed @ induct_prem_prem_thms)),
49377
7003b063a985 use right version of "mk_UnIN"
blanchet
parents: 49376
diff changeset
   130
     rtac (mk_UnIN pp jj),
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   131
     mk_induct_prem_prem_endgame_tac ctxt qq]) (rev ixs)) 1;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   132
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   133
fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k ixs =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   134
  EVERY [mk_induct_prepare_prem_tac n m k,
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   135
    mk_induct_prepare_prem_prems_tac (length ixs), atac 1,
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   136
    mk_induct_discharge_prem_prems_tac ctxt nn ixs set_natural's pre_set_defs];
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   137
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   138
fun mk_induct_tac ctxt ns mss ixsss ctr_defs fld_induct' set_natural's pre_set_defss =
49376
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   139
  let
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   140
    val nn = length ns;
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   141
    val n = Integer.sum ns;
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   142
  in
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   143
    mk_induct_prelude_tac ctxt ctr_defs fld_induct' THEN
49376
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   144
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
49378
19237e465055 fixed issue with bound variables in prem prems + tuning
blanchet
parents: 49377
diff changeset
   145
      pre_set_defss mss (unflat mss (1 upto n)) ixsss)
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   146
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   147
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   148
end;