src/HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
changeset 49590 43e2d0e10876
parent 49589 71aa74965bc9
child 49591 91b228e26348
equal deleted inserted replaced
49589:71aa74965bc9 49590:43e2d0e10876
    10   val sum_prod_thms_map: thm list
    10   val sum_prod_thms_map: thm list
    11   val sum_prod_thms_set: thm list
    11   val sum_prod_thms_set: thm list
    12   val sum_prod_thms_rel: thm list
    12   val sum_prod_thms_rel: thm list
    13 
    13 
    14   val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
    14   val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
       
    15   val mk_coinduct_tac: Proof.context -> int -> int list -> thm -> thm list -> thm list ->
       
    16     thm list -> thm list list -> thm list list list -> thm list list -> tactic
    15   val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
    17   val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
    16   val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
    18   val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
    17     tactic
    19     tactic
    18   val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
    20   val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
    19   val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
    21   val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
    20   val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
    22   val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
    21   val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
    23   val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
    22     thm -> thm list -> thm list list -> tactic
    24     thm list -> thm -> thm list -> thm list list -> tactic
    23   val mk_inject_tac: Proof.context -> thm -> thm -> tactic
    25   val mk_inject_tac: Proof.context -> thm -> thm -> tactic
    24   val mk_rec_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
    26   val mk_rec_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
    25 end;
    27 end;
    26 
    28 
    27 structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
    29 structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
    29 
    31 
    30 open BNF_Tactics
    32 open BNF_Tactics
    31 open BNF_Util
    33 open BNF_Util
    32 open BNF_FP
    34 open BNF_FP
    33 
    35 
       
    36 val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
       
    37 val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
       
    38 
    34 val sum_prod_thms_map = @{thms id_apply map_pair_simp sum_map.simps prod.cases};
    39 val sum_prod_thms_map = @{thms id_apply map_pair_simp sum_map.simps prod.cases};
    35 
       
    36 val sum_prod_thms_set0 =
    40 val sum_prod_thms_set0 =
    37   @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
    41   @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
    38       Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
    42       Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
    39       mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
    43       mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
    40 
       
    41 val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
    44 val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
    42 
       
    43 val sum_prod_thms_rel = @{thms prod.cases prod_rel_def sum.cases sum_rel_def};
    45 val sum_prod_thms_rel = @{thms prod.cases prod_rel_def sum.cases sum_rel_def};
    44 
    46 
    45 fun mk_proj n k T = funpow n (fn t => Abs (Name.uu, T, t)) (Bound (n - k));
    47 val ss_if_True_False = ss_only @{thms if_True if_False};
    46 
    48 
    47 fun inst_as_projs ctxt n k thm =
    49 fun mk_proj T k =
       
    50   let val binders = binder_types T in
       
    51     fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (length binders - k))
       
    52   end;
       
    53 
       
    54 fun inst_as_projs ctxt k thm =
    48   let
    55   let
    49     val fs =
    56     val fs =
    50       Term.add_vars (prop_of thm) []
    57       Term.add_vars (prop_of thm) []
    51       |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
    58       |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
    52     val cfs =
    59     val cfs =
    53       map (fn f as (_, T) => (certify ctxt (Var f), certify ctxt (mk_proj n k (domain_type T)))) fs;
    60       map (fn f as (_, T) => (certify ctxt (Var f), certify ctxt (mk_proj T k))) fs;
    54   in
    61   in
    55     Drule.cterm_instantiate cfs thm
    62     Drule.cterm_instantiate cfs thm
    56   end;
    63   end;
    57 
    64 
    58 val inst_as_projs_tac = PRIMITIVE ooo inst_as_projs;
    65 val inst_as_projs_tac = PRIMITIVE oo inst_as_projs;
    59 
    66 
    60 fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
    67 fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
    61   unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
    68   unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
    62   (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
    69   (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
    63    REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
    70    REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
    91 
    98 
    92 fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
    99 fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
    93   unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
   100   unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
    94     rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
   101     rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
    95 
   102 
    96 val corec_like_ss = ss_only @{thms if_True if_False};
       
    97 
       
    98 fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
   103 fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
    99   unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
   104   unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
   100   subst_tac ctxt NONE [ctor_dtor_corec_like] 1 THEN asm_simp_tac corec_like_ss 1 THEN
   105   subst_tac ctxt NONE [ctor_dtor_corec_like] 1 THEN asm_simp_tac ss_if_True_False 1 THEN
   101   unfold_thms_tac ctxt (pre_map_def :: sum_prod_thms_map @ map_ids) THEN
   106   unfold_thms_tac ctxt (pre_map_def :: sum_prod_thms_map @ map_ids) THEN
   102   unfold_thms_tac ctxt @{thms id_def} THEN
   107   unfold_thms_tac ctxt @{thms id_def} THEN
   103   TRY ((rtac refl ORELSE' subst_tac ctxt NONE @{thms unit_eq} THEN' rtac refl) 1);
   108   TRY ((rtac refl ORELSE' subst_tac ctxt NONE @{thms unit_eq} THEN' rtac refl) 1);
   104 
   109 
   105 fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
   110 fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
   106   EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
   111   EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
   107       case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
   112       case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
   108       asm_simp_tac (ss_only @{thms simp_thms(7,8,12,14,22,24)}) 1 THEN
   113       asm_simp_tac (ss_only basic_simp_thms) 1 THEN
   109       (if is_refl disc then all_tac else rtac disc 1))
   114       (if is_refl disc then all_tac else rtac disc 1))
   110     (map rtac case_splits' @ [K all_tac]) corec_likes discs);
   115     (map rtac case_splits' @ [K all_tac]) corec_likes discs);
   111 
   116 
   112 val solve_prem_prem_tac =
   117 val solve_prem_prem_tac =
   113   REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
   118   REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
   127         (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
   132         (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
   128       if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
   133       if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
   129       mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
   134       mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
   130   end;
   135   end;
   131 
   136 
   132 fun mk_induct_tac ctxt ns mss kkss ctr_defs ctor_induct' set_natural's pre_set_defss =
   137 fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' set_natural's pre_set_defss =
   133   let
   138   let val n = Integer.sum ns in
   134     val nn = length ns;
   139     unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_as_projs_tac ctxt 1 THEN
   135     val n = Integer.sum ns;
       
   136   in
       
   137     unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_as_projs_tac ctxt 1 1 THEN
       
   138     EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
   140     EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
   139       pre_set_defss mss (unflat mss (1 upto n)) kkss)
   141       pre_set_defss mss (unflat mss (1 upto n)) kkss)
   140   end;
   142   end;
   141 
   143 
       
   144 fun mk_coinduct_choose_prem_tac nn kk =
       
   145   EVERY' [rtac allI, rtac allI, rtac impI,
       
   146     select_prem_tac nn (dtac meta_spec) kk,
       
   147     dtac meta_spec, dtac meta_mp, atac];
       
   148 
       
   149 fun mk_coinduct_same_ctr ctxt pre_rel_def dtor_ctor ctr_def sels disc =
       
   150   hyp_subst_tac THEN' subst_tac ctxt (SOME [1, 2]) [ctr_def] THEN'
       
   151   SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN'
       
   152   (rtac refl ORELSE' atac ORELSE'
       
   153    full_simp_tac (ss_only (disc :: more_simp_thms)));
       
   154 
       
   155 fun mk_coinduct_distinct_ctrs discs =
       
   156   hyp_subst_tac THEN' full_simp_tac (ss_only (refl :: discs @ basic_simp_thms));
       
   157 
       
   158 fun mk_coinduct_exhaust_tac ctxt exhaust k i =
       
   159   rtac exhaust i THEN inst_as_projs_tac ctxt k;
       
   160 
       
   161 fun mk_coinduct_discharge_prem_tac ctxt nn kk n pre_rel_def dtor_ctor exhaust ctr_defs selss discs =
       
   162   let val ks = 1 upto n in
       
   163     EVERY' ([mk_coinduct_choose_prem_tac nn kk, mk_coinduct_exhaust_tac ctxt exhaust 1,
       
   164         hyp_subst_tac] @
       
   165       map4 (fn k => fn ctr_def => fn sels => fn disc =>
       
   166         EVERY' (mk_coinduct_exhaust_tac ctxt exhaust 2 ::
       
   167           map2 (fn k' => fn disc' =>
       
   168             if k' = k then mk_coinduct_same_ctr ctxt pre_rel_def dtor_ctor ctr_def sels disc
       
   169             else mk_coinduct_distinct_ctrs [disc, disc']) ks discs)) ks ctr_defs selss discs)
       
   170   end;
       
   171 
       
   172 fun mk_coinduct_tac ctxt nn ns dtor_coinduct pre_rel_defs dtor_ctors exhausts ctr_defss selsss
       
   173     discss =
       
   174   (rtac dtor_coinduct THEN'
       
   175    EVERY' (map8 (mk_coinduct_discharge_prem_tac ctxt nn)
       
   176      (1 upto nn) ns pre_rel_defs dtor_ctors exhausts ctr_defss selsss discss)) 1
       
   177 
   142 end;
   178 end;