src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML
author blanchet
Tue, 28 May 2013 08:36:12 +0200
changeset 52195 056ec8201667
parent 52194 6289b167bbab
child 52214 4cc5a80bba80
permissions -rw-r--r--
clean up list of theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     3
    Copyright   2012
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     4
49389
blanchet
parents: 49385
diff changeset
     5
Tactics for datatype and codatatype sugar.
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     6
*)
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     7
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
     8
signature BNF_FP_DEF_SUGAR_TACTICS =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     9
sig
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    10
  val sum_prod_thms_map: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    11
  val sum_prod_thms_set: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    12
  val sum_prod_thms_rel: thm list
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    13
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    14
  val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    15
  val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    16
    thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> tactic
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
    17
  val mk_coiter_tac: thm list -> thm list -> thm list -> thm list -> thm list -> thm -> thm ->
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
    18
    thm -> Proof.context -> tactic
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    19
  val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    20
    tactic
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
    21
  val mk_disc_coiter_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    22
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    23
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    24
  val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    25
    thm list -> thm -> thm list -> thm list list -> tactic
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
    26
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
    27
  val mk_iter_tac: thm list -> thm list -> thm list -> thm list -> thm -> thm -> Proof.context
49670
c7a034d01936 changed the type of the recursor for nested recursion
blanchet
parents: 49668
diff changeset
    28
    -> tactic
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    29
end;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    30
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
    31
structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    32
struct
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    33
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    34
open BNF_Tactics
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    35
open BNF_Util
51850
106afdf5806c renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
blanchet
parents: 51843
diff changeset
    36
open BNF_FP_Util
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    37
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    38
val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    39
val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    40
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
    41
val sum_prod_thms_map = @{thms id_apply map_pair_simp prod.cases sum.cases sum_map.simps};
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    42
val sum_prod_thms_set0 =
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    43
  @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    44
      Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    45
      mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    46
val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51850
diff changeset
    47
val sum_prod_thms_rel = @{thms prod_rel_simp sum_rel_simps id_apply};
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    48
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
    49
val ss_if_True_False = simpset_of (ss_only @{thms if_True if_False} @{context});
49589
71aa74965bc9 generalized tactic a bit
blanchet
parents: 49586
diff changeset
    50
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    51
fun mk_proj T k =
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    52
  let val binders = binder_types T in
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    53
    fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (length binders - k))
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    54
  end;
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    55
49668
blanchet
parents: 49665
diff changeset
    56
fun hhf_concl_conv cv ctxt ct =
blanchet
parents: 49665
diff changeset
    57
  (case Thm.term_of ct of
blanchet
parents: 49665
diff changeset
    58
    Const (@{const_name all}, _) $ Abs _ =>
blanchet
parents: 49665
diff changeset
    59
    Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct
blanchet
parents: 49665
diff changeset
    60
  | _ => Conv.concl_conv ~1 cv ct);
blanchet
parents: 49665
diff changeset
    61
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    62
fun inst_as_projs ctxt k thm =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    63
  let
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    64
    val fs =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    65
      Term.add_vars (prop_of thm) []
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    66
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    67
    val cfs =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    68
      map (fn f as (_, T) => (certify ctxt (Var f), certify ctxt (mk_proj T k))) fs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    69
  in
49384
94ad5ba23541 took out one rotate_tac
blanchet
parents: 49383
diff changeset
    70
    Drule.cterm_instantiate cfs thm
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    71
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    72
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    73
val inst_as_projs_tac = PRIMITIVE oo inst_as_projs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    74
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    75
fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    76
  unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
49264
9059e0dbdbc1 implemented and use "mk_sum_casesN_balanced"
blanchet
parents: 49263
diff changeset
    77
  (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    78
   REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    79
   rtac refl) 1;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49127
diff changeset
    80
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    81
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    82
  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN rtac sumEN' 1 THEN
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    83
  unfold_thms_tac ctxt @{thms all_prod_eq} THEN
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    84
  EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    85
    etac meta_mp, atac]) (1 upto n)) 1;
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    86
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    87
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    88
  (rtac iffI THEN'
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    89
   EVERY' (map3 (fn cTs => fn cx => fn th =>
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    90
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    91
     SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    92
     atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])) 1;
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    93
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    94
fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    95
  unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    96
  rtac @{thm sum.distinct(1)} 1;
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    97
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    98
fun mk_inject_tac ctxt ctr_def ctor_inject =
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
    99
  unfold_thms_tac ctxt [ctr_def] THEN rtac (ctor_inject RS ssubst) 1 THEN
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49501
diff changeset
   100
  unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   101
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   102
val iter_unfold_thms =
52195
056ec8201667 clean up list of theorems
blanchet
parents: 52194
diff changeset
   103
  @{thms comp_def convol_def fst_conv id_def prod_case_Pair_iden snd_conv
056ec8201667 clean up list of theorems
blanchet
parents: 52194
diff changeset
   104
      split_conv unit_case_Unity} @ sum_prod_thms_map;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   105
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   106
fun mk_iter_tac pre_map_defs map_comp's map_ids'' iter_defs ctor_iter ctr_def ctxt =
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   107
  unfold_thms_tac ctxt (ctr_def :: ctor_iter :: iter_defs @ pre_map_defs @ map_comp's @
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   108
    map_ids'' @ iter_unfold_thms) THEN rtac refl 1;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   109
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   110
val coiter_unfold_thms =
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
   111
  @{thms id_def ident_o_ident sum_case_if sum_case_o_inj} @ sum_prod_thms_map;
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
   112
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   113
fun mk_coiter_tac coiter_defs map_comps'' map_comp's map_ids'' map_if_distribs
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   114
    ctor_dtor_coiter pre_map_def ctr_def ctxt =
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   115
  unfold_thms_tac ctxt (ctr_def :: coiter_defs) THEN
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   116
  (rtac (ctor_dtor_coiter RS trans) THEN'
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   117
    asm_simp_tac (put_simpset ss_if_True_False ctxt)) 1 THEN_MAYBE
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
   118
  (unfold_thms_tac ctxt (pre_map_def :: map_comp's @ map_comps'' @ map_ids'' @ map_if_distribs @
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   119
    coiter_unfold_thms) THEN
49680
00290dc6bfad made (co)rec tactic more robust when the simplifier succeeds early
blanchet
parents: 49672
diff changeset
   120
   (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)) 1);
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
   121
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   122
fun mk_disc_coiter_iff_tac case_splits' coiters discs ctxt =
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   123
  EVERY (map3 (fn case_split_tac => fn coiter_thm => fn disc =>
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   124
      case_split_tac 1 THEN unfold_thms_tac ctxt [coiter_thm] THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   125
      asm_simp_tac (ss_only basic_simp_thms ctxt) 1 THEN
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49482
diff changeset
   126
      (if is_refl disc then all_tac else rtac disc 1))
51843
899663644482 rationalized terminology (iterator = fold or rec, xxfoo = (co)foo or (un)foo)
blanchet
parents: 51798
diff changeset
   127
    (map rtac case_splits' @ [K all_tac]) coiters discs);
49482
e6d6869eed08 generate coiter_iff and corec_iff theorems
blanchet
parents: 49463
diff changeset
   128
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   129
fun solve_prem_prem_tac ctxt =
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   130
  REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   131
    hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   132
  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
49426
6d05b8452cf3 got rid of one "auto" in induction tactic
blanchet
parents: 49391
diff changeset
   133
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   134
fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_map's pre_set_defs =
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   135
  EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   136
     SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_map's @ sum_prod_thms_set0)),
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   137
     solve_prem_prem_tac ctxt]) (rev kks)) 1;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   138
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   139
fun mk_induct_discharge_prem_tac ctxt nn n set_map's pre_set_defs m k kks =
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   140
  let val r = length kks in
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   141
    EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   142
      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   143
    EVERY [REPEAT_DETERM_N r
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   144
        (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   145
      if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   146
      mk_induct_leverage_prem_prems_tac ctxt nn kks set_map's pre_set_defs]
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   147
  end;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   148
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   149
fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' set_map's pre_set_defss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   150
  let val n = Integer.sum ns in
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   151
    unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_as_projs_tac ctxt 1 THEN
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   152
    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_map's) pre_set_defss
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51717
diff changeset
   153
      mss (unflat mss (1 upto n)) kkss)
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   154
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   155
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   156
fun mk_coinduct_same_ctr ctxt rel_eqs pre_rel_def dtor_ctor ctr_def discs sels =
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   157
  hyp_subst_tac ctxt THEN'
49665
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   158
  CONVERSION (hhf_concl_conv
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   159
    (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'
49642
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49639
diff changeset
   160
  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   161
  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN'
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   162
  (atac ORELSE' REPEAT o etac conjE THEN'
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   163
     full_simp_tac
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   164
       (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN_MAYBE'
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   165
     REPEAT o hyp_subst_tac ctxt THEN' REPEAT o rtac conjI THEN' REPEAT o rtac refl);
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   166
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   167
fun mk_coinduct_distinct_ctrs ctxt discs discs' =
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   168
  hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   169
  full_simp_tac (ss_only (refl :: no_refl (discs @ discs') @ basic_simp_thms) ctxt);
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   170
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   171
fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   172
    discss selss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   173
  let val ks = 1 upto n in
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   174
    EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk, dtac
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   175
        meta_spec, dtac meta_mp, atac, rtac exhaust, K (inst_as_projs_tac ctxt 1),
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   176
        hyp_subst_tac ctxt] @
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   177
      map4 (fn k => fn ctr_def => fn discs => fn sels =>
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   178
        EVERY' ([rtac exhaust, K (inst_as_projs_tac ctxt 2)] @
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   179
          map2 (fn k' => fn discs' =>
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   180
            if k' = k then
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   181
              mk_coinduct_same_ctr ctxt rel_eqs' pre_rel_def dtor_ctor ctr_def discs sels
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   182
            else
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49683
diff changeset
   183
              mk_coinduct_distinct_ctrs ctxt discs discs') ks discss)) ks ctr_defs discss selss)
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   184
  end;
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   185
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   186
fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs dtor_ctors exhausts ctr_defss
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   187
    discsss selsss =
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   188
  (rtac dtor_coinduct' THEN'
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   189
   EVERY' (map8 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   190
     (1 upto nn) ns pre_rel_defs dtor_ctors exhausts ctr_defss discsss selsss)) 1;
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   191
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   192
end;