src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Mon, 17 Sep 2012 21:13:30 +0200
changeset 50444 64ac3471005a
parent 50443 93f158d59ead
child 50445 6df729c6a1a6
permissions -rw-r--r--
cleaner way of dealing with the set functions of sums and products
blanchet@50138
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
blanchet@50138
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50138
     3
    Copyright   2012
blanchet@50138
     4
blanchet@50404
     5
Tactics for datatype and codatatype sugar.
blanchet@50138
     6
*)
blanchet@50138
     7
blanchet@50138
     8
signature BNF_FP_SUGAR_TACTICS =
blanchet@50138
     9
sig
blanchet@50145
    10
  val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
blanchet@50241
    11
  val mk_coiter_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
blanchet@50176
    12
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
blanchet@50387
    13
  val mk_fld_iff_unf_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
blanchet@50387
    14
    tactic
blanchet@50142
    15
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
blanchet@50444
    16
  val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
blanchet@50444
    17
    thm -> thm list -> thm list list -> tactic
blanchet@50141
    18
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
blanchet@50241
    19
  val mk_iter_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
blanchet@50138
    20
end;
blanchet@50138
    21
blanchet@50138
    22
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
blanchet@50138
    23
struct
blanchet@50138
    24
blanchet@50140
    25
open BNF_Tactics
blanchet@50138
    26
open BNF_Util
blanchet@50279
    27
open BNF_FP_Util
blanchet@50138
    28
blanchet@50383
    29
val meta_mp = @{thm meta_mp};
blanchet@50383
    30
val meta_spec = @{thm meta_spec};
blanchet@50383
    31
blanchet@50406
    32
fun inst_spurious_fs lthy thm =
blanchet@50383
    33
  let
blanchet@50399
    34
    val fs =
blanchet@50383
    35
      Term.add_vars (prop_of thm) []
blanchet@50383
    36
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
blanchet@50399
    37
    val cfs =
blanchet@50399
    38
      map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
blanchet@50383
    39
  in
blanchet@50399
    40
    Drule.cterm_instantiate cfs thm
blanchet@50383
    41
  end;
blanchet@50383
    42
blanchet@50406
    43
val inst_spurious_fs_tac = PRIMITIVE o inst_spurious_fs;
blanchet@50383
    44
blanchet@50145
    45
fun mk_case_tac ctxt n k m case_def ctr_def unf_fld =
blanchet@50145
    46
  Local_Defs.unfold_tac ctxt [case_def, ctr_def, unf_fld] THEN
blanchet@50279
    47
  (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
blanchet@50145
    48
   REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
blanchet@50145
    49
   rtac refl) 1;
blanchet@50145
    50
blanchet@50176
    51
fun mk_exhaust_tac ctxt n ctr_defs fld_iff_unf sumEN' =
blanchet@50278
    52
  Local_Defs.unfold_tac ctxt (fld_iff_unf :: ctr_defs) THEN rtac sumEN' 1 THEN
blanchet@50140
    53
  Local_Defs.unfold_tac ctxt @{thms all_prod_eq} THEN
blanchet@50383
    54
  EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
blanchet@50383
    55
    etac meta_mp, atac]) (1 upto n)) 1;
blanchet@50140
    56
blanchet@50138
    57
fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
blanchet@50138
    58
  (rtac iffI THEN'
blanchet@50138
    59
   EVERY' (map3 (fn cTs => fn cx => fn th =>
blanchet@50138
    60
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
blanchet@50138
    61
     SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
blanchet@50138
    62
     atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
blanchet@50138
    63
blanchet@50142
    64
fun mk_half_distinct_tac ctxt fld_inject ctr_defs =
blanchet@50142
    65
  Local_Defs.unfold_tac ctxt (fld_inject :: @{thms sum.inject} @ ctr_defs) THEN
blanchet@50142
    66
  rtac @{thm sum.distinct(1)} 1;
blanchet@50142
    67
blanchet@50141
    68
fun mk_inject_tac ctxt ctr_def fld_inject =
blanchet@50150
    69
  Local_Defs.unfold_tac ctxt [ctr_def] THEN rtac (fld_inject RS ssubst) 1 THEN
blanchet@50150
    70
  Local_Defs.unfold_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
blanchet@50141
    71
blanchet@50226
    72
val iter_like_thms =
blanchet@50248
    73
  @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
blanchet@50248
    74
      split_conv};
blanchet@50220
    75
blanchet@50244
    76
fun mk_iter_like_tac pre_map_defs map_ids iter_like_defs fld_iter_like ctr_def ctxt =
blanchet@50244
    77
  Local_Defs.unfold_tac ctxt (ctr_def :: fld_iter_like :: iter_like_defs @ pre_map_defs @ map_ids @
blanchet@50247
    78
    iter_like_thms) THEN Local_Defs.unfold_tac ctxt @{thms id_def} THEN rtac refl 1;
blanchet@50220
    79
blanchet@50228
    80
val coiter_like_ss = ss_only @{thms if_True if_False};
blanchet@50291
    81
val coiter_like_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
blanchet@50228
    82
blanchet@50244
    83
fun mk_coiter_like_tac coiter_like_defs map_ids fld_unf_coiter_like pre_map_def ctr_def ctxt =
blanchet@50228
    84
  Local_Defs.unfold_tac ctxt (ctr_def :: coiter_like_defs) THEN
blanchet@50241
    85
  subst_tac ctxt [fld_unf_coiter_like] 1 THEN asm_simp_tac coiter_like_ss 1 THEN
blanchet@50247
    86
  Local_Defs.unfold_tac ctxt (pre_map_def :: coiter_like_thms @ map_ids) THEN
blanchet@50250
    87
  Local_Defs.unfold_tac ctxt @{thms id_def} THEN
blanchet@50377
    88
  TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
blanchet@50228
    89
blanchet@50442
    90
val maybe_singletonI_tac = atac ORELSE' rtac @{thm singletonI};
blanchet@50383
    91
blanchet@50444
    92
val solve_prem_prem_tac =
blanchet@50444
    93
  REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
blanchet@50444
    94
    hyp_subst_tac ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
blanchet@50444
    95
  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
blanchet@50441
    96
blanchet@50383
    97
val induct_prem_prem_thms =
blanchet@50441
    98
  @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex_eq_empty UN_compreh_bex_eq_singleton
blanchet@50444
    99
      UN_insert Un_assoc Un_empty_left Un_empty_right Un_iff Union_Un_distrib collect_def[abs_def]
blanchet@50444
   100
      eq_UN_compreh_Un fst_conv image_def o_apply snd_conv snd_prod_fun sum.cases sup_bot_right
blanchet@50444
   101
      fst_map_pair map_pair_simp mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps
blanchet@50444
   102
      sum_set_simps};
blanchet@50383
   103
blanchet@50444
   104
fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs =
blanchet@50444
   105
  EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
blanchet@50404
   106
     SELECT_GOAL (Local_Defs.unfold_tac ctxt
blanchet@50404
   107
       (pre_set_defs @ set_natural's @ induct_prem_prem_thms)),
blanchet@50444
   108
     solve_prem_prem_tac]) (rev kks)) 1;
blanchet@50393
   109
blanchet@50444
   110
fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k kks =
blanchet@50444
   111
  let val r = length kks in
blanchet@50406
   112
    EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
blanchet@50406
   113
      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
blanchet@50406
   114
    EVERY [REPEAT_DETERM_N r
blanchet@50406
   115
        (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
blanchet@50406
   116
      if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
blanchet@50444
   117
      mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
blanchet@50406
   118
  end;
blanchet@50383
   119
blanchet@50444
   120
fun mk_induct_tac ctxt ns mss kkss ctr_defs fld_induct' set_natural's pre_set_defss =
blanchet@50391
   121
  let
blanchet@50391
   122
    val nn = length ns;
blanchet@50391
   123
    val n = Integer.sum ns;
blanchet@50391
   124
  in
blanchet@50406
   125
    Local_Defs.unfold_tac ctxt ctr_defs THEN rtac fld_induct' 1 THEN inst_spurious_fs_tac ctxt THEN
blanchet@50391
   126
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
blanchet@50444
   127
      pre_set_defss mss (unflat mss (1 upto n)) kkss)
blanchet@50383
   128
  end;
blanchet@50383
   129
blanchet@50138
   130
end;