src/HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Wed, 26 Sep 2012 10:00:59 +0200
changeset 49586 d5e342ffe91e
parent 49585 5c4a12550491
child 49589 71aa74965bc9
permissions -rw-r--r--
parameterized "subst_tac"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49504
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
49123
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
49389
blanchet
parents: 49385
diff changeset
     5
Tactics for datatype and codatatype sugar.
49123
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
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    10
  val sum_prod_thms_map: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    11
  val sum_prod_thms_set: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    12
  val sum_prod_thms_rel: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    13
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    14
  val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    15
  val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    16
  val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    17
    tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    18
  val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    19
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    20
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    21
  val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    22
    thm -> thm list -> thm list list -> tactic
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    23
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    24
  val mk_rec_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
    25
end;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    26
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    27
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    28
struct
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    29
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    30
open BNF_Tactics
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    31
open BNF_Util
49457
1d2825673cec renamed "bnf_fp_util.ML" to "bnf_fp.ML"
blanchet
parents: 49436
diff changeset
    32
open BNF_FP
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    33
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    34
val sum_prod_thms_map = @{thms id_apply map_pair_simp sum_map.simps prod.cases};
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    35
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    36
val sum_prod_thms_set0 =
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    37
  @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    38
      Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    39
      mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    40
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    41
val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    42
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    43
val sum_prod_thms_rel = @{thms prod.cases prod_rel_def sum.cases sum_rel_def};
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    44
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
    45
fun inst_spurious_fs lthy thm =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    46
  let
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    47
    val fs =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    48
      Term.add_vars (prop_of thm) []
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    49
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    50
    val cfs =
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    51
      map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    52
  in
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    53
    Drule.cterm_instantiate cfs thm
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    54
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    55
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
    56
val inst_spurious_fs_tac = PRIMITIVE o inst_spurious_fs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    57
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    58
fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    59
  unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
49264
9059e0dbdbc1 implemented and use "mk_sum_casesN_balanced"
blanchet
parents: 49263
diff changeset
    60
  (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
    61
   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
    62
   rtac refl) 1;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    63
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    64
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    65
  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN rtac sumEN' 1 THEN
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    66
  unfold_thms_tac ctxt @{thms all_prod_eq} THEN
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    67
  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
    68
    etac meta_mp, atac]) (1 upto n)) 1;
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    69
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    70
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    71
  (rtac iffI THEN'
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    72
   EVERY' (map3 (fn cTs => fn cx => fn th =>
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    73
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    74
     SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    75
     atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])) 1;
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    76
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    77
fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    78
  unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    79
  rtac @{thm sum.distinct(1)} 1;
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    80
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    81
fun mk_inject_tac ctxt ctr_def ctor_inject =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    82
  unfold_thms_tac ctxt [ctr_def] THEN rtac (ctor_inject RS ssubst) 1 THEN
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    83
  unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    84
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    85
(*TODO: Try "sum_prod_thms_map" here, enriched with a few theorems*)
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    86
val rec_like_unfold_thms =
49539
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    87
  @{thms comp_def convol_def id_apply map_pair_def prod_case_Pair_iden sum.simps(5,6) sum_map.simps
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    88
      split_conv unit_case_Unity};
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
    89
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    90
fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    91
  unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    92
    rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
    93
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    94
val corec_like_ss = ss_only @{thms if_True if_False};
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
    95
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    96
fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    97
  unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
49586
d5e342ffe91e parameterized "subst_tac"
blanchet
parents: 49585
diff changeset
    98
  subst_tac ctxt NONE [ctor_dtor_corec_like] 1 THEN asm_simp_tac corec_like_ss 1 THEN
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    99
  unfold_thms_tac ctxt (pre_map_def :: sum_prod_thms_map @ map_ids) THEN
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   100
  unfold_thms_tac ctxt @{thms id_def} THEN
49586
d5e342ffe91e parameterized "subst_tac"
blanchet
parents: 49585
diff changeset
   101
  TRY ((rtac refl ORELSE' subst_tac ctxt NONE @{thms unit_eq} THEN' rtac refl) 1);
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
   102
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   103
fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   104
  EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   105
      case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
49482
e6d6869eed08 generate coiter_iff and corec_iff theorems
blanchet
parents: 49463
diff changeset
   106
      asm_simp_tac (ss_only @{thms simp_thms(7,8,12,14,22,24)}) 1 THEN
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49482
diff changeset
   107
      (if is_refl disc then all_tac else rtac disc 1))
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   108
    (map rtac case_splits' @ [K all_tac]) corec_likes discs);
49482
e6d6869eed08 generate coiter_iff and corec_iff theorems
blanchet
parents: 49463
diff changeset
   109
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   110
val solve_prem_prem_tac =
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   111
  REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   112
    hyp_subst_tac ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   113
  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
49426
6d05b8452cf3 got rid of one "auto" in induction tactic
blanchet
parents: 49391
diff changeset
   114
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   115
fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs =
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   116
  EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
   117
     SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_natural's @ sum_prod_thms_set0)),
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   118
     solve_prem_prem_tac]) (rev kks)) 1;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   119
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   120
fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k kks =
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   121
  let val r = length kks in
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   122
    EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   123
      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   124
    EVERY [REPEAT_DETERM_N r
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   125
        (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   126
      if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   127
      mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   128
  end;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   129
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
   130
fun mk_induct_tac ctxt ns mss kkss ctr_defs ctor_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
   131
  let
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   132
    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
   133
    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
   134
  in
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   135
    unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_spurious_fs_tac ctxt THEN
49376
c6366fd0415a select the right premise in "mk_induct_discharge_prem_prems_tac" instead of relying on backtracking
blanchet
parents: 49375
diff changeset
   136
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   137
      pre_set_defss mss (unflat mss (1 upto n)) kkss)
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   138
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   139
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   140
end;