src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML
author wenzelm
Sun, 10 Aug 2014 14:34:43 +0200
changeset 57882 38bf4de248a6
parent 57815 f97643a56615
parent 57824 615223745d4e
child 57891 d23a85b59dec
permissions -rw-r--r--
merged -- with manual conflict resolution for src/HOL/SMT_Examples/SMT_Examples.certs2, src/HOL/SMT_Examples/SMT_Word_Examples.certs2, src/Doc/Prog_Prove/document/intro-isabelle.tex;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/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
57668
blanchet
parents: 57563
diff changeset
     3
    Author:     Martin Desharnais, TU Muenchen
blanchet
parents: 57563
diff changeset
     4
    Copyright   2012, 2013, 2014
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     5
49389
blanchet
parents: 49385
diff changeset
     6
Tactics for datatype and codatatype sugar.
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     7
*)
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
     8
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
     9
signature BNF_FP_DEF_SUGAR_TACTICS =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    10
sig
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    11
  val sumprod_thms_map: thm list
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    12
  val sumprod_thms_set: thm list
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    13
  val sumprod_thms_rel: thm list
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    14
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 ->
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    16
    thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    17
    thm list list list -> tactic
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
    18
  val mk_corec_tac: thm list -> thm list -> thm -> thm -> thm -> 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
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
    21
  val mk_disc_corec_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
    22
  val mk_disc_map_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    23
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    24
  val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    25
  val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    26
    thm list -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    27
  val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
    28
  val mk_rec_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context ->
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    29
    tactic
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
    30
  val mk_rel_cases_tac:  Proof.context -> cterm -> cterm -> thm -> thm list -> thm list ->
57824
615223745d4e made tactic more robust w.r.t. dead variables; tuned;
desharna
parents: 57670
diff changeset
    31
    thm list -> thm list -> thm list -> tactic
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
    32
  val mk_rel_coinduct0_tac: Proof.context -> thm -> cterm list -> thm list -> thm list ->
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
    33
    thm list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
    34
    thm list -> thm list -> thm list -> tactic
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
    35
  val mk_rel_induct0_tac: Proof.context -> thm -> thm list -> cterm list -> thm list ->
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
    36
    thm list list -> thm list -> thm list -> thm list -> thm list -> tactic
57563
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
    37
  val mk_rel_sel_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
    38
    thm list -> thm list -> tactic
57046
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
    39
  val mk_sel_map_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
    40
  val mk_sel_set_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
56956
7425fa3763ff generate 'set_empty' theorem for BNFs
desharna
parents: 56765
diff changeset
    41
  val mk_set_empty_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
    42
  val mk_set_induct0_tac: Proof.context -> cterm list -> thm list -> thm list -> thm list ->
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
    43
    thm list -> thm list -> thm list -> thm list -> tactic
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    44
end;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    45
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
    46
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
    47
struct
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    48
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
    49
open Ctr_Sugar_Util
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    50
open BNF_Tactics
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    51
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
    52
open BNF_FP_Util
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    53
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    54
val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    55
val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
57303
498a62e65f5f tune the implementation of 'rel_coinduct'
desharna
parents: 57301
diff changeset
    56
val simp_thms' = @{thms simp_thms(6,7,8,11,12,15,16,22,24)};
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    57
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    58
val sumprod_thms_map = @{thms id_apply map_prod_simp prod.case sum.case map_sum.simps};
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    59
val sumprod_thms_set =
57815
f97643a56615 generate nicer 'set' theorems for (co)datatypes
blanchet
parents: 57700
diff changeset
    60
  @{thms UN_empty UN_insert UN_simps(10) UN_iff Un_empty_left Un_empty_right Un_iff Union_Un_distrib
f97643a56615 generate nicer 'set' theorems for (co)datatypes
blanchet
parents: 57700
diff changeset
    61
      image_iff o_apply map_prod_simp mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps};
56765
644f0d4820a1 cleaner 'rel_inject' theorems
blanchet
parents: 56245
diff changeset
    62
val sumprod_thms_rel = @{thms rel_sum_simps rel_prod_apply prod.inject id_apply conj_assoc};
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    63
49668
blanchet
parents: 49665
diff changeset
    64
fun hhf_concl_conv cv ctxt ct =
blanchet
parents: 49665
diff changeset
    65
  (case Thm.term_of ct of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55966
diff changeset
    66
    Const (@{const_name Pure.all}, _) $ Abs _ =>
49668
blanchet
parents: 49665
diff changeset
    67
    Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct
blanchet
parents: 49665
diff changeset
    68
  | _ => Conv.concl_conv ~1 cv ct);
blanchet
parents: 49665
diff changeset
    69
54922
blanchet
parents: 54837
diff changeset
    70
fun co_induct_inst_as_projs ctxt k thm =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    71
  let
54922
blanchet
parents: 54837
diff changeset
    72
    val fs = Term.add_vars (prop_of thm) []
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    73
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
    74
    fun mk_cfp (f as (_, T)) =
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
    75
      (certify ctxt (Var f), certify ctxt (mk_proj T (num_binder_types T) k));
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
    76
    val cfps = map mk_cfp fs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    77
  in
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
    78
    Drule.cterm_instantiate cfps thm
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    79
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    80
54922
blanchet
parents: 54837
diff changeset
    81
val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    82
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    83
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
52324
blanchet
parents: 52323
diff changeset
    84
  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac sumEN') THEN
blanchet
parents: 52323
diff changeset
    85
  HEADGOAL (EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k,
blanchet
parents: 52323
diff changeset
    86
    REPEAT_DETERM o dtac meta_spec, etac meta_mp, atac]) (1 upto n)));
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    87
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    88
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
52324
blanchet
parents: 52323
diff changeset
    89
  HEADGOAL (rtac iffI THEN'
blanchet
parents: 52323
diff changeset
    90
    EVERY' (map3 (fn cTs => fn cx => fn th =>
blanchet
parents: 52323
diff changeset
    91
      dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
blanchet
parents: 52323
diff changeset
    92
      SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
blanchet
parents: 52323
diff changeset
    93
      atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    94
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    95
fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    96
  unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
52324
blanchet
parents: 52323
diff changeset
    97
  HEADGOAL (rtac @{thm sum.distinct(1)});
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
    98
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
    99
fun mk_inject_tac ctxt ctr_def ctor_inject abs_inject =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   100
  unfold_thms_tac ctxt [ctr_def] THEN
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   101
  HEADGOAL (rtac (ctor_inject RS ssubst)) THEN
56765
644f0d4820a1 cleaner 'rel_inject' theorems
blanchet
parents: 56245
diff changeset
   102
  unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject prod.inject conj_assoc}) THEN
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   103
  HEADGOAL (rtac refl);
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   104
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   105
val rec_unfold_thms =
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
   106
  @{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
   107
      case_unit_Unity} @ sumprod_thms_map;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   108
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   109
fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec fp_abs_inverse abs_inverse ctr_def ctxt =
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   110
  unfold_thms_tac ctxt (ctr_def :: ctor_rec :: fp_abs_inverse :: abs_inverse :: rec_defs @
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   111
    pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac refl);
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   112
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
   113
val corec_unfold_thms = @{thms id_def} @ sumprod_thms_map;
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
   114
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   115
fun mk_corec_tac corec_defs map_ident0s ctor_dtor_corec pre_map_def abs_inverse ctr_def ctxt =
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   116
  let
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   117
    val ss = ss_only (pre_map_def :: abs_inverse :: map_ident0s @ corec_unfold_thms @
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   118
      @{thms o_apply vimage2p_def if_True if_False}) ctxt;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   119
  in
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   120
    unfold_thms_tac ctxt (ctr_def :: corec_defs) THEN
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   121
    HEADGOAL (rtac (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   122
    HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong))
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   123
  end;
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
   124
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   125
fun mk_disc_corec_iff_tac case_splits' corecs discs ctxt =
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   126
  EVERY (map3 (fn case_split_tac => fn corec_thm => fn disc =>
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   127
      HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN
52324
blanchet
parents: 52323
diff changeset
   128
      HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN
blanchet
parents: 52323
diff changeset
   129
      (if is_refl disc then all_tac else HEADGOAL (rtac disc)))
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   130
    (map rtac case_splits' @ [K all_tac]) corecs discs);
49482
e6d6869eed08 generate coiter_iff and corec_iff theorems
blanchet
parents: 49463
diff changeset
   131
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
   132
fun mk_disc_map_iff_tac ctxt ct exhaust discs maps =
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
   133
  TRYALL Goal.conjunction_tac THEN
57558
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   134
  ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   135
    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   136
  unfold_thms_tac ctxt maps THEN
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   137
  unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   138
    handle THM _ => thm RS eqTrueI) discs) THEN
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   139
  ALLGOALS (rtac refl ORELSE' rtac TrueI);
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
   140
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   141
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
   142
  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
   143
    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
   144
  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
49426
6d05b8452cf3 got rid of one "auto" in induction tactic
blanchet
parents: 49391
diff changeset
   145
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   146
fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   147
    pre_set_defs =
52324
blanchet
parents: 52323
diff changeset
   148
  HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   149
    SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
   150
      sumprod_thms_set)),
52324
blanchet
parents: 52323
diff changeset
   151
    solve_prem_prem_tac ctxt]) (rev kks)));
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   152
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   153
fun mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps pre_set_defs m k
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   154
    kks =
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   155
  let val r = length kks in
52324
blanchet
parents: 52323
diff changeset
   156
    HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
blanchet
parents: 52323
diff changeset
   157
      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   158
    EVERY [REPEAT_DETERM_N r
52324
blanchet
parents: 52323
diff changeset
   159
        (HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54241
diff changeset
   160
      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac,
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   161
      mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   162
        pre_set_defs]
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   163
  end;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   164
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   165
fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' fp_abs_inverses abs_inverses set_maps
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   166
    pre_set_defss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   167
  let val n = Integer.sum ns in
54922
blanchet
parents: 54837
diff changeset
   168
    unfold_thms_tac ctxt ctr_defs THEN HEADGOAL (rtac ctor_induct') THEN
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
   169
    co_induct_inst_as_projs_tac ctxt 0 THEN
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   170
    EVERY (map4 (EVERY oooo map3 o
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   171
        mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   172
      pre_set_defss mss (unflat mss (1 upto n)) kkss)
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   173
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   174
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   175
fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def fp_abs_inverse abs_inverse dtor_ctor ctr_def
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   176
    discs sels =
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   177
  hyp_subst_tac ctxt THEN'
49665
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   178
  CONVERSION (hhf_concl_conv
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   179
    (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
   180
  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   181
  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor ::
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
   182
    sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN'
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   183
  (atac ORELSE' REPEAT o etac conjE THEN'
56765
644f0d4820a1 cleaner 'rel_inject' theorems
blanchet
parents: 56245
diff changeset
   184
     full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
54241
357988ad95ec strengthened tactic
blanchet
parents: 54198
diff changeset
   185
     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
357988ad95ec strengthened tactic
blanchet
parents: 54198
diff changeset
   186
     REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   187
52966
blanchet
parents: 52659
diff changeset
   188
fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
54198
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   189
  let
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   190
    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   191
      |> distinct Thm.eq_thm_prop;
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   192
  in
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   193
    hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   194
    full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   195
  end;
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   196
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   197
fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   198
    dtor_ctor exhaust ctr_defs discss selss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   199
  let val ks = 1 upto n in
54837
5bc637eb60c0 tuning whitespace
blanchet
parents: 54742
diff changeset
   200
    EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk,
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
   201
        dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   202
        hyp_subst_tac ctxt] @
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   203
      map4 (fn k => fn ctr_def => fn discs => fn sels =>
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54922
diff changeset
   204
        EVERY' ([rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   205
          map2 (fn k' => fn discs' =>
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   206
            if k' = k then
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   207
              mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def fp_abs_inverse abs_inverse
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   208
                dtor_ctor ctr_def discs sels
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   209
            else
52966
blanchet
parents: 52659
diff changeset
   210
              mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   211
  end;
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   212
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   213
fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses abs_inverses
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   214
    dtor_ctors exhausts ctr_defss discsss selsss =
52324
blanchet
parents: 52323
diff changeset
   215
  HEADGOAL (rtac dtor_coinduct' THEN'
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   216
    EVERY' (map10 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   217
      (1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   218
      selsss));
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   219
57824
615223745d4e made tactic more robust w.r.t. dead variables; tuned;
desharna
parents: 57670
diff changeset
   220
fun mk_rel_cases_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts rel_eqs=
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   221
  HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   222
    rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   223
      hyp_subst_tac ctxt) THEN
57824
615223745d4e made tactic more robust w.r.t. dead variables; tuned;
desharna
parents: 57670
diff changeset
   224
  unfold_thms_tac ctxt (rel_eqs @ injects @ rel_injects @ @{thms conj_imp_eq_imp_imp simp_thms(6)
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   225
    True_implies_equals conj_imp_eq_imp_imp} @
57529
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   226
    map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   227
    map (fn thm => thm RS eqTrueI) rel_injects) THEN
57562
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   228
  TRYALL (atac ORELSE' etac FalseE ORELSE'
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   229
    (REPEAT_DETERM o dtac @{thm meta_spec} THEN'
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   230
     TRY o filter_prems_tac
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   231
       (forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN'
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   232
     REPEAT_DETERM o (dtac @{thm meta_mp} THEN' rtac refl) THEN' Goal.assume_rule_tac ctxt));
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   233
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   234
fun mk_rel_coinduct0_tac ctxt dtor_rel_coinduct cts assms exhausts discss selss ctor_defss
57558
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   235
    dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses nesting_rel_eqs =
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   236
  rtac dtor_rel_coinduct 1 THEN
57670
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   237
   EVERY (map11 (fn ct => fn assm => fn exhaust => fn discs => fn sels => fn ctor_defs =>
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   238
     fn dtor_ctor => fn ctor_inject => fn abs_inject => fn rel_pre_def => fn abs_inverse =>
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   239
      (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   240
         (dtac (rotate_prems (~1) (cterm_instantiate_pos [NONE, NONE, NONE, NONE, SOME ct]
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   241
            @{thm arg_cong2} RS iffD1)) THEN'
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   242
          atac THEN' atac THEN' hyp_subst_tac ctxt THEN' dtac assm THEN'
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   243
          REPEAT_DETERM o etac conjE))) 1 THEN
57668
blanchet
parents: 57563
diff changeset
   244
      unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN
57528
9afc832252a3 refactor some tactics
desharna
parents: 57525
diff changeset
   245
      unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject ::
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   246
        abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @ @{thms BNF_Comp.id_bnf_comp_def
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   247
        rel_sum_simps rel_prod_apply vimage2p_def Inl_Inr_False iffD2[OF eq_False Inr_not_Inl]
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   248
        sum.inject prod.inject}) THEN
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   249
      REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac conjE) THEN' (REPEAT_DETERM o rtac conjI) THEN'
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   250
        (rtac refl ORELSE' atac))))
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   251
    cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   252
      abs_inverses);
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   253
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   254
fun mk_rel_induct0_tac ctxt ctor_rel_induct assms cterms exhausts ctor_defss ctor_injects
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   255
    rel_pre_list_defs Abs_inverses nesting_rel_eqs =
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   256
  rtac ctor_rel_induct 1 THEN EVERY (map6 (fn cterm => fn exhaust => fn ctor_defs =>
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   257
      fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse =>
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   258
        HEADGOAL (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   259
          (rtac (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2)
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   260
            THEN' atac THEN' atac THEN' TRY o resolve_tac assms))) THEN
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   261
        unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   262
          @{thms BNF_Comp.id_bnf_comp_def vimage2p_def}) THEN
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   263
        TRYALL (hyp_subst_tac ctxt) THEN
57824
615223745d4e made tactic more robust w.r.t. dead variables; tuned;
desharna
parents: 57670
diff changeset
   264
        unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_apply Inl_Inr_False
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   265
          Inr_Inl_False  sum.inject prod.inject}) THEN
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   266
        TRYALL (etac FalseE ORELSE' (REPEAT_DETERM o etac conjE) THEN' atac))
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   267
    cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses);
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   268
57563
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   269
fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts =
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   270
  HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   271
    rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   272
      hyp_subst_tac ctxt) THEN
57824
615223745d4e made tactic more robust w.r.t. dead variables; tuned;
desharna
parents: 57670
diff changeset
   273
  unfold_thms_tac ctxt (sels @ rel_injects @ @{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @
57563
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   274
    ((discs @ distincts) RL @{thms iffD2[OF eq_True] iffD2[OF eq_False]}) @
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   275
    (rel_injects RL @{thms iffD2[OF eq_True]}) @
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   276
    (rel_distincts RL @{thms iffD2[OF eq_False]})) THEN
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   277
  TRYALL (resolve_tac [TrueI, refl]);
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   278
57046
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   279
fun mk_sel_map_tac ctxt ct exhaust discs maps sels =
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   280
  TRYALL Goal.conjunction_tac THEN
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   281
    ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   282
      REPEAT_DETERM o hyp_subst_tac ctxt) THEN
57529
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   283
    unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
57046
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   284
      @{thms not_True_eq_False not_False_eq_True}) THEN
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   285
    TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN
57528
9afc832252a3 refactor some tactics
desharna
parents: 57525
diff changeset
   286
    unfold_thms_tac ctxt (maps @ sels) THEN
57046
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   287
    ALLGOALS (rtac refl);
b3613d7e108b generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents: 56991
diff changeset
   288
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   289
fun mk_sel_set_tac ctxt ct exhaust discs sels sets =
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   290
  TRYALL Goal.conjunction_tac THEN
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   291
    ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   292
      REPEAT_DETERM o hyp_subst_tac ctxt) THEN
57529
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   293
    unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   294
      @{thms not_True_eq_False not_False_eq_True}) THEN
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   295
    TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN
57528
9afc832252a3 refactor some tactics
desharna
parents: 57525
diff changeset
   296
    unfold_thms_tac ctxt (sels @ sets) THEN
57558
6bb3dd7f8097 took out 'rel_cases' for now because of failing tactic
blanchet
parents: 57529
diff changeset
   297
    ALLGOALS (REPEAT o (resolve_tac @{thms UnI1 UnI2 imageI} ORELSE'
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   298
        eresolve_tac @{thms UN_I UN_I[rotated] imageE} ORELSE'
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   299
        hyp_subst_tac ctxt) THEN'
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   300
      (rtac @{thm singletonI} ORELSE' atac));
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   301
56956
7425fa3763ff generate 'set_empty' theorem for BNFs
desharna
parents: 56765
diff changeset
   302
fun mk_set_empty_tac ctxt ct exhaust sets discs =
7425fa3763ff generate 'set_empty' theorem for BNFs
desharna
parents: 56765
diff changeset
   303
  TRYALL Goal.conjunction_tac THEN
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
   304
  ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
56990
299b026cc5af fix 'set_empty' theorem when the discriminator is 'op ='
desharna
parents: 56956
diff changeset
   305
    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
299b026cc5af fix 'set_empty' theorem when the discriminator is 'op ='
desharna
parents: 56956
diff changeset
   306
  unfold_thms_tac ctxt (sets @ map_filter (fn thm =>
57529
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   307
    SOME (thm RS eqFalseI) handle THM _ => NONE) discs) THEN
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
   308
  ALLGOALS (rtac refl ORELSE' etac FalseE);
56956
7425fa3763ff generate 'set_empty' theorem for BNFs
desharna
parents: 56765
diff changeset
   309
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   310
fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   311
    Abs_pre_inverses =
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   312
  let
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   313
    val assms_ctor_defs =
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   314
      map (unfold_thms ctxt (@{thm BNF_Comp.id_bnf_comp_def} :: ctor_defs)) assms;
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   315
    val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   316
      |> map2 (fn ct => cterm_instantiate_pos [NONE, SOME ct]) cts;
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   317
  in
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   318
    ALLGOALS (resolve_tac dtor_set_inducts) THEN
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   319
    TRYALL (resolve_tac exhausts' THEN_ALL_NEW
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   320
      (resolve_tac (map (fn ct => refl RS
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   321
         cterm_instantiate_pos (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts)
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   322
        THEN' atac THEN' hyp_subst_tac ctxt)) THEN
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   323
    unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   324
      @{thms BNF_Comp.id_bnf_comp_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   325
        Un_empty_left Un_empty_right empty_iff singleton_iff}) THEN
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   326
    REPEAT_DETERM (HEADGOAL
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   327
      (TRY o etac UnE THEN' TRY o etac @{thm singletonE} THEN' TRY o hyp_subst_tac ctxt THEN'
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   328
       REPEAT_DETERM o eresolve_tac @{thms UN_E UnE singletonE} THEN'
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   329
       fold (curry (op ORELSE')) (map (fn thm =>
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   330
         funpow (length (prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms_ctor_defs)
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   331
         (etac FalseE)))
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   332
  end;
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   333
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   334
end;