src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Fri, 14 Sep 2012 22:23:11 +0200
changeset 50400 83b867707af2
parent 50399 94ad5ba23541
child 50404 da621dc65146
permissions -rw-r--r--
merged two unfold steps
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@50138
     5
Tactics for the LFP/GFP 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@50391
    16
  val mk_induct_tac: Proof.context -> int list -> int list list ->
blanchet@50393
    17
    ((int * int) * (int * int)) list list list -> thm list -> thm -> thm list -> thm list list ->
blanchet@50393
    18
    tactic
blanchet@50141
    19
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
blanchet@50241
    20
  val mk_iter_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
blanchet@50138
    21
end;
blanchet@50138
    22
blanchet@50138
    23
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
blanchet@50138
    24
struct
blanchet@50138
    25
blanchet@50140
    26
open BNF_Tactics
blanchet@50138
    27
open BNF_Util
blanchet@50279
    28
open BNF_FP_Util
blanchet@50138
    29
blanchet@50383
    30
val meta_mp = @{thm meta_mp};
blanchet@50383
    31
val meta_spec = @{thm meta_spec};
blanchet@50383
    32
blanchet@50385
    33
fun smash_spurious_fs lthy thm =
blanchet@50383
    34
  let
blanchet@50399
    35
    val fs =
blanchet@50383
    36
      Term.add_vars (prop_of thm) []
blanchet@50383
    37
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
blanchet@50399
    38
    val cfs =
blanchet@50399
    39
      map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
blanchet@50383
    40
  in
blanchet@50399
    41
    Drule.cterm_instantiate cfs thm
blanchet@50383
    42
  end;
blanchet@50383
    43
blanchet@50385
    44
val smash_spurious_fs_tac = PRIMITIVE o smash_spurious_fs;
blanchet@50383
    45
blanchet@50145
    46
fun mk_case_tac ctxt n k m case_def ctr_def unf_fld =
blanchet@50145
    47
  Local_Defs.unfold_tac ctxt [case_def, ctr_def, unf_fld] THEN
blanchet@50279
    48
  (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
blanchet@50145
    49
   REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
blanchet@50145
    50
   rtac refl) 1;
blanchet@50145
    51
blanchet@50176
    52
fun mk_exhaust_tac ctxt n ctr_defs fld_iff_unf sumEN' =
blanchet@50278
    53
  Local_Defs.unfold_tac ctxt (fld_iff_unf :: ctr_defs) THEN rtac sumEN' 1 THEN
blanchet@50140
    54
  Local_Defs.unfold_tac ctxt @{thms all_prod_eq} THEN
blanchet@50383
    55
  EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
blanchet@50383
    56
    etac meta_mp, atac]) (1 upto n)) 1;
blanchet@50140
    57
blanchet@50138
    58
fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
blanchet@50138
    59
  (rtac iffI THEN'
blanchet@50138
    60
   EVERY' (map3 (fn cTs => fn cx => fn th =>
blanchet@50138
    61
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
blanchet@50138
    62
     SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
blanchet@50138
    63
     atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
blanchet@50138
    64
blanchet@50142
    65
fun mk_half_distinct_tac ctxt fld_inject ctr_defs =
blanchet@50142
    66
  Local_Defs.unfold_tac ctxt (fld_inject :: @{thms sum.inject} @ ctr_defs) THEN
blanchet@50142
    67
  rtac @{thm sum.distinct(1)} 1;
blanchet@50142
    68
blanchet@50141
    69
fun mk_inject_tac ctxt ctr_def fld_inject =
blanchet@50150
    70
  Local_Defs.unfold_tac ctxt [ctr_def] THEN rtac (fld_inject RS ssubst) 1 THEN
blanchet@50150
    71
  Local_Defs.unfold_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
blanchet@50141
    72
blanchet@50226
    73
val iter_like_thms =
blanchet@50248
    74
  @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
blanchet@50248
    75
      split_conv};
blanchet@50220
    76
blanchet@50244
    77
fun mk_iter_like_tac pre_map_defs map_ids iter_like_defs fld_iter_like ctr_def ctxt =
blanchet@50244
    78
  Local_Defs.unfold_tac ctxt (ctr_def :: fld_iter_like :: iter_like_defs @ pre_map_defs @ map_ids @
blanchet@50247
    79
    iter_like_thms) THEN Local_Defs.unfold_tac ctxt @{thms id_def} THEN rtac refl 1;
blanchet@50220
    80
blanchet@50228
    81
val coiter_like_ss = ss_only @{thms if_True if_False};
blanchet@50291
    82
val coiter_like_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
blanchet@50228
    83
blanchet@50244
    84
fun mk_coiter_like_tac coiter_like_defs map_ids fld_unf_coiter_like pre_map_def ctr_def ctxt =
blanchet@50228
    85
  Local_Defs.unfold_tac ctxt (ctr_def :: coiter_like_defs) THEN
blanchet@50241
    86
  subst_tac ctxt [fld_unf_coiter_like] 1 THEN asm_simp_tac coiter_like_ss 1 THEN
blanchet@50247
    87
  Local_Defs.unfold_tac ctxt (pre_map_def :: coiter_like_thms @ map_ids) THEN
blanchet@50250
    88
  Local_Defs.unfold_tac ctxt @{thms id_def} THEN
blanchet@50377
    89
  TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
blanchet@50228
    90
blanchet@50383
    91
fun mk_induct_prelude_tac ctxt ctr_defs fld_induct' =
blanchet@50385
    92
  Local_Defs.unfold_tac ctxt ctr_defs THEN rtac fld_induct' 1 THEN smash_spurious_fs_tac ctxt;
blanchet@50383
    93
blanchet@50383
    94
fun mk_induct_prepare_prem_tac n m k =
blanchet@50383
    95
  EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
blanchet@50383
    96
    REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1;
blanchet@50383
    97
blanchet@50395
    98
(* FIXME: why not in "Pure"? *)
blanchet@50395
    99
fun prefer_tac i = defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1);
blanchet@50395
   100
blanchet@50395
   101
fun mk_induct_prepare_prem_prems_tac r =
blanchet@50395
   102
  REPEAT_DETERM_N r (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2) THEN
blanchet@50395
   103
  PRIMITIVE Raw_Simplifier.norm_hhf;
blanchet@50383
   104
blanchet@50383
   105
val induct_prem_prem_thms =
blanchet@50383
   106
  @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex UN_insert Un_assoc[symmetric] Un_empty_left
blanchet@50387
   107
      Un_empty_right Union_Un_distrib collect_def[abs_def] fst_conv image_def o_apply snd_conv
blanchet@50387
   108
      snd_prod_fun sum.cases sup_bot_right fst_map_pair map_pair_simp sum_map.simps};
blanchet@50383
   109
blanchet@50387
   110
(* These rules interfere with the "set_natural'" properties of "sum" and "prod", so we explicitly
blanchet@50387
   111
   delay them. *)
blanchet@50387
   112
val induct_prem_prem_thms_delayed =
blanchet@50387
   113
  @{thms fsts_def[abs_def] snds_def[abs_def] sum_setl_def[abs_def] sum_setr_def[abs_def]};
blanchet@50387
   114
blanchet@50399
   115
(* TODO: Get rid of "auto_tac" (or at least use a naked context) *)
blanchet@50394
   116
fun mk_induct_prem_prem_endgame_tac _ 0 = atac ORELSE' rtac @{thm singletonI}
blanchet@50394
   117
  | mk_induct_prem_prem_endgame_tac ctxt qq =
blanchet@50394
   118
    REPEAT_DETERM_N qq o
blanchet@50394
   119
      (SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms Union_iff bex_simps(6)}) THEN'
blanchet@50394
   120
       etac @{thm induct_set_step}) THEN'
blanchet@50398
   121
    atac ORELSE' SELECT_GOAL (auto_tac ctxt);
blanchet@50393
   122
blanchet@50393
   123
fun mk_induct_discharge_prem_prems_tac ctxt nn ixs set_natural's pre_set_defs =
blanchet@50393
   124
  EVERY' (maps (fn ((pp, jj), (qq, kk)) =>
blanchet@50399
   125
      [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
blanchet@50400
   126
       SELECT_GOAL (Local_Defs.unfold_tac ctxt
blanchet@50400
   127
         (pre_set_defs @ set_natural's @ induct_prem_prem_thms)),
blanchet@50397
   128
       SELECT_GOAL (Local_Defs.unfold_tac ctxt
blanchet@50397
   129
         (induct_prem_prem_thms_delayed @ induct_prem_prem_thms)),
blanchet@50398
   130
       rtac (mk_UnIN pp jj) THEN' mk_induct_prem_prem_endgame_tac ctxt qq ORELSE'
blanchet@50398
   131
         SELECT_GOAL (auto_tac ctxt)])
blanchet@50397
   132
    (rev ixs)) 1;
blanchet@50383
   133
blanchet@50393
   134
fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k ixs =
blanchet@50383
   135
  EVERY [mk_induct_prepare_prem_tac n m k,
blanchet@50393
   136
    mk_induct_prepare_prem_prems_tac (length ixs), atac 1,
blanchet@50393
   137
    mk_induct_discharge_prem_prems_tac ctxt nn ixs set_natural's pre_set_defs];
blanchet@50383
   138
blanchet@50393
   139
fun mk_induct_tac ctxt ns mss ixsss ctr_defs fld_induct' set_natural's pre_set_defss =
blanchet@50391
   140
  let
blanchet@50391
   141
    val nn = length ns;
blanchet@50391
   142
    val n = Integer.sum ns;
blanchet@50391
   143
  in
blanchet@50383
   144
    mk_induct_prelude_tac ctxt ctr_defs fld_induct' THEN
blanchet@50391
   145
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
blanchet@50393
   146
      pre_set_defss mss (unflat mss (1 upto n)) ixsss)
blanchet@50383
   147
  end;
blanchet@50383
   148
blanchet@50138
   149
end;