src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Fri, 14 Sep 2012 22:23:10 +0200
changeset 50390 993677c1cf2a
parent 50387 c62165188971
child 50391 c6366fd0415a
permissions -rw-r--r--
tuned code before fixing "mk_induct_discharge_prem_prems_tac"
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@50390
    16
  val mk_induct_tac: Proof.context -> int list -> int list list -> (int * int) list list list ->
blanchet@50390
    17
    thm list -> 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@50385
    32
fun smash_spurious_fs lthy thm =
blanchet@50383
    33
  let
blanchet@50383
    34
    val spurious_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@50383
    37
    val cxs =
blanchet@50383
    38
      map (fn s as (_, T) =>
blanchet@50383
    39
        (certify lthy (Var s), certify lthy (id_abs (domain_type T)))) spurious_fs;
blanchet@50383
    40
  in
blanchet@50383
    41
    Drule.cterm_instantiate cxs 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@50387
    98
fun mk_induct_prepare_prem_prems_tac 0 = all_tac
blanchet@50387
    99
  | mk_induct_prepare_prem_prems_tac r =
blanchet@50387
   100
    REPEAT_DETERM_N r ((rotate_tac ~1) 1 THEN dtac meta_mp 1 THEN
blanchet@50383
   101
      defer_tac 2 THEN PRIMITIVE (Thm.permute_prems 0 ~1) THEN rotate_tac 1 1) THEN
blanchet@50383
   102
    PRIMITIVE Raw_Simplifier.norm_hhf;
blanchet@50383
   103
blanchet@50383
   104
val induct_prem_prem_thms =
blanchet@50383
   105
  @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex UN_insert Un_assoc[symmetric] Un_empty_left
blanchet@50387
   106
      Un_empty_right Union_Un_distrib collect_def[abs_def] fst_conv image_def o_apply snd_conv
blanchet@50387
   107
      snd_prod_fun sum.cases sup_bot_right fst_map_pair map_pair_simp sum_map.simps};
blanchet@50383
   108
blanchet@50387
   109
(* These rules interfere with the "set_natural'" properties of "sum" and "prod", so we explicitly
blanchet@50387
   110
   delay them. *)
blanchet@50387
   111
val induct_prem_prem_thms_delayed =
blanchet@50387
   112
  @{thms fsts_def[abs_def] snds_def[abs_def] sum_setl_def[abs_def] sum_setr_def[abs_def]};
blanchet@50387
   113
blanchet@50387
   114
(* TODO: Get rid of the "blast_tac" *)
blanchet@50390
   115
fun mk_induct_discharge_prem_prems_tac ctxt ppis set_natural's pre_set_defs =
blanchet@50390
   116
  EVERY' (maps (fn (pp, i) =>
blanchet@50390
   117
    [(* ### select_prem_tac pp (dtac meta_spec) i, *) dtac meta_spec, rotate_tac ~1, etac meta_mp,
blanchet@50390
   118
     SELECT_GOAL (Local_Defs.unfold_tac ctxt pre_set_defs), (* ### why on a line of its own? *)
blanchet@50387
   119
     SELECT_GOAL (Local_Defs.unfold_tac ctxt (set_natural's @ induct_prem_prem_thms)),
blanchet@50383
   120
     SELECT_GOAL (Local_Defs.unfold_tac ctxt
blanchet@50387
   121
       (induct_prem_prem_thms_delayed @ induct_prem_prem_thms)),
blanchet@50390
   122
     TRY o rtac (mk_UnIN pp i), (*#####*)
blanchet@50383
   123
     atac ORELSE'
blanchet@50383
   124
     rtac @{thm singletonI} ORELSE'
blanchet@50387
   125
     (REPEAT_DETERM o (SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms Union_iff bex_simps(6)}) THEN'
blanchet@50387
   126
          etac @{thm induct_set_step}) THEN'
blanchet@50390
   127
      (atac ORELSE' blast_tac ctxt))]) (rev ppis)) 1;
blanchet@50383
   128
blanchet@50390
   129
fun mk_induct_discharge_prem_tac ctxt n set_natural's pre_set_defs m k ppis =
blanchet@50383
   130
  EVERY [mk_induct_prepare_prem_tac n m k,
blanchet@50390
   131
    mk_induct_prepare_prem_prems_tac (length ppis), atac 1,
blanchet@50390
   132
    mk_induct_discharge_prem_prems_tac ctxt ppis set_natural's pre_set_defs];
blanchet@50383
   133
blanchet@50390
   134
fun mk_induct_tac ctxt ns mss ppisss ctr_defs fld_induct' set_natural's pre_set_defss =
blanchet@50387
   135
  let val n = Integer.sum ns in
blanchet@50383
   136
    mk_induct_prelude_tac ctxt ctr_defs fld_induct' THEN
blanchet@50390
   137
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt n set_natural's)
blanchet@50390
   138
      pre_set_defss mss (unflat mss (1 upto n)) ppisss)
blanchet@50383
   139
  end;
blanchet@50383
   140
blanchet@50138
   141
end;