src/HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
changeset 49510 ba50d204095e
parent 49509 163914705f8d
child 49539 be6cbf960aa7
equal deleted inserted replaced
49509:163914705f8d 49510:ba50d204095e
       
     1 (*  Title:      HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
       
     2     Author:     Jasmin Blanchette, TU Muenchen
       
     3     Copyright   2012
       
     4 
       
     5 Tactics for datatype and codatatype sugar.
       
     6 *)
       
     7 
       
     8 signature BNF_FP_SUGAR_TACTICS =
       
     9 sig
       
    10   val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
       
    11   val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
       
    12   val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
       
    13     tactic
       
    14   val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
       
    15   val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
       
    16   val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
       
    17   val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
       
    18     thm -> thm list -> thm list list -> tactic
       
    19   val mk_inject_tac: Proof.context -> thm -> thm -> tactic
       
    20   val mk_rec_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
       
    21 end;
       
    22 
       
    23 structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
       
    24 struct
       
    25 
       
    26 open BNF_Tactics
       
    27 open BNF_Util
       
    28 open BNF_FP
       
    29 
       
    30 val meta_mp = @{thm meta_mp};
       
    31 val meta_spec = @{thm meta_spec};
       
    32 
       
    33 fun inst_spurious_fs lthy thm =
       
    34   let
       
    35     val fs =
       
    36       Term.add_vars (prop_of thm) []
       
    37       |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
       
    38     val cfs =
       
    39       map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
       
    40   in
       
    41     Drule.cterm_instantiate cfs thm
       
    42   end;
       
    43 
       
    44 val inst_spurious_fs_tac = PRIMITIVE o inst_spurious_fs;
       
    45 
       
    46 fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
       
    47   unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
       
    48   (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
       
    49    REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
       
    50    rtac refl) 1;
       
    51 
       
    52 fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
       
    53   unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN rtac sumEN' 1 THEN
       
    54   unfold_thms_tac ctxt @{thms all_prod_eq} THEN
       
    55   EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
       
    56     etac meta_mp, atac]) (1 upto n)) 1;
       
    57 
       
    58 fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
       
    59   (rtac iffI THEN'
       
    60    EVERY' (map3 (fn cTs => fn cx => fn th =>
       
    61      dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
       
    62      SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
       
    63      atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])) 1;
       
    64 
       
    65 fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
       
    66   unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
       
    67   rtac @{thm sum.distinct(1)} 1;
       
    68 
       
    69 fun mk_inject_tac ctxt ctr_def ctor_inject =
       
    70   unfold_thms_tac ctxt [ctr_def] THEN rtac (ctor_inject RS ssubst) 1 THEN
       
    71   unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
       
    72 
       
    73 val rec_like_unfold_thms =
       
    74   @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
       
    75       split_conv};
       
    76 
       
    77 fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
       
    78   unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
       
    79     rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
       
    80 
       
    81 val corec_like_ss = ss_only @{thms if_True if_False};
       
    82 val corec_like_unfold_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
       
    83 
       
    84 fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
       
    85   unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
       
    86   subst_tac ctxt [ctor_dtor_corec_like] 1 THEN asm_simp_tac corec_like_ss 1 THEN
       
    87   unfold_thms_tac ctxt (pre_map_def :: corec_like_unfold_thms @ map_ids) THEN
       
    88   unfold_thms_tac ctxt @{thms id_def} THEN
       
    89   TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
       
    90 
       
    91 fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
       
    92   EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
       
    93       case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
       
    94       asm_simp_tac (ss_only @{thms simp_thms(7,8,12,14,22,24)}) 1 THEN
       
    95       (if is_refl disc then all_tac else rtac disc 1))
       
    96     (map rtac case_splits' @ [K all_tac]) corec_likes discs);
       
    97 
       
    98 val solve_prem_prem_tac =
       
    99   REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
       
   100     hyp_subst_tac ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
       
   101   (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
       
   102 
       
   103 val induct_prem_prem_thms =
       
   104   @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
       
   105       Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
       
   106       mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
       
   107 
       
   108 fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs =
       
   109   EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
       
   110      SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_natural's @ induct_prem_prem_thms)),
       
   111      solve_prem_prem_tac]) (rev kks)) 1;
       
   112 
       
   113 fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k kks =
       
   114   let val r = length kks in
       
   115     EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
       
   116       REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
       
   117     EVERY [REPEAT_DETERM_N r
       
   118         (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
       
   119       if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
       
   120       mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
       
   121   end;
       
   122 
       
   123 fun mk_induct_tac ctxt ns mss kkss ctr_defs ctor_induct' set_natural's pre_set_defss =
       
   124   let
       
   125     val nn = length ns;
       
   126     val n = Integer.sum ns;
       
   127   in
       
   128     unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_spurious_fs_tac ctxt THEN
       
   129     EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
       
   130       pre_set_defss mss (unflat mss (1 upto n)) kkss)
       
   131   end;
       
   132 
       
   133 end;