src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML
author blanchet
Sun, 11 Sep 2016 15:37:09 +0200
changeset 63842 f738df816abf
parent 63841 813a574da746
child 63845 61a03e429cbd
permissions -rw-r--r--
strengthened tactics
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
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
    11
  val sumprod_thms_rel: thm list
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    12
  val sumprod_thms_set: thm list
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
    13
  val basic_sumprod_thms_set: thm list
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49542
diff changeset
    14
58093
6f37a300c82b generate 'case_transfer' for (co)datatypes
desharna
parents: 58044
diff changeset
    15
  val mk_case_transfer_tac: Proof.context -> thm -> thm list -> tactic
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    16
  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
    17
    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
    18
    thm list list list -> tactic
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
    19
  val mk_corec_tac: thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context -> tactic
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
    20
  val mk_corec_disc_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
58734
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
    21
  val mk_co_rec_o_map_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> thm -> thm ->
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
    22
    thm Seq.seq
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
    23
  val mk_corec_transfer_tac: Proof.context -> cterm list -> cterm list -> thm list -> thm list ->
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
    24
    thm list -> thm list -> thm list -> ''a list -> ''a list list -> ''a list list list list ->
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
    25
    ''a list list list list -> tactic
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
    26
  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
    27
    tactic
58327
a147bd03cad0 make 'ctr_transfer' tactic more robust
desharna
parents: 58326
diff changeset
    28
  val mk_ctr_transfer_tac: Proof.context -> thm list -> thm list -> tactic
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
    29
  val mk_disc_transfer_tac: Proof.context -> thm -> thm -> thm list -> tactic
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49160
diff changeset
    30
  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
    31
  val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    32
  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
    33
    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
    34
  val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
    35
  val mk_map_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> thm list -> tactic
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
    36
  val mk_map_disc_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
58326
7e142efcee1a make 'rel_sel' and 'map_sel' tactics more robust
desharna
parents: 58181
diff changeset
    37
  val mk_map_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list ->
7e142efcee1a make 'rel_sel' and 'map_sel' tactics more robust
desharna
parents: 58181
diff changeset
    38
    thm list -> tactic
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
    39
  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
    40
    tactic
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
    41
  val mk_rec_transfer_tac: Proof.context -> int -> int list -> cterm list -> cterm list ->
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
    42
    term list list list list -> thm list -> thm list -> thm list -> thm list -> tactic
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
    43
  val mk_rel_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> thm list -> tactic
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
    44
  val mk_rel_case_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
    45
    thm list -> thm list -> tactic
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
    46
  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
    47
    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
    48
    thm list -> thm list -> thm list -> tactic
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
    49
  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
    50
    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
    51
  val mk_rel_sel_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->
58326
7e142efcee1a make 'rel_sel' and 'map_sel' tactics more robust
desharna
parents: 58181
diff changeset
    52
    thm list -> thm list -> thm list -> tactic
58676
cdf84b6e1297 generate 'sel_transfer' for (co)datatypes
desharna
parents: 58634
diff changeset
    53
  val mk_sel_transfer_tac: Proof.context -> int -> thm list -> thm -> tactic
57893
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
    54
  val mk_set_cases_tac: Proof.context -> cterm -> thm list -> thm -> thm list -> tactic
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
    55
  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
    56
    thm list -> thm list -> thm list -> thm list -> tactic
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
    57
  val mk_set_intros_tac: Proof.context -> thm list -> tactic
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
    58
  val mk_set_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    59
end;
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    60
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
    61
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
    62
struct
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    63
56991
8e9ca31e9b8e generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents: 56990
diff changeset
    64
open Ctr_Sugar_Util
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
    65
open BNF_Tactics
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    66
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
    67
open BNF_FP_Util
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
    68
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
    69
val case_sum_transfer = @{thm case_sum_transfer};
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63068
diff changeset
    70
val case_sum_transfer_eq = @{thm case_sum_transfer[of "op =" _ "op =", simplified sum.rel_eq]};
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
    71
val case_prod_transfer = @{thm case_prod_transfer};
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63068
diff changeset
    72
val case_prod_transfer_eq = @{thm case_prod_transfer[of "op =" "op =", simplified prod.rel_eq]};
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
    73
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    74
val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
    75
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
    76
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
    77
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
    78
val sumprod_thms_map = @{thms id_apply map_prod_simp prod.case sum.case map_sum.simps};
63841
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
    79
val sumprod_thms_rel = @{thms rel_sum_simps rel_prod_inject prod.inject id_apply conj_assoc};
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
    80
val basic_sumprod_thms_set =
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
    81
  @{thms UN_empty UN_insert  UN_iff Un_empty_left Un_empty_right Un_iff Union_Un_distrib
63835
blanchet
parents: 63170
diff changeset
    82
      o_apply map_prod_simp mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps};
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
    83
val sumprod_thms_set = @{thms UN_simps(10) image_iff} @ basic_sumprod_thms_set;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    84
58359
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
    85
fun is_def_looping def =
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
    86
  (case Thm.prop_of def of
63835
blanchet
parents: 63170
diff changeset
    87
    Const (@{const_name Pure.eq}, _) $ lhs $ rhs => Term.exists_subterm (curry (op aconv) lhs) rhs
58359
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
    88
  | _ => false);
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
    89
49668
blanchet
parents: 49665
diff changeset
    90
fun hhf_concl_conv cv ctxt ct =
blanchet
parents: 49665
diff changeset
    91
  (case Thm.term_of ct of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55966
diff changeset
    92
    Const (@{const_name Pure.all}, _) $ Abs _ =>
49668
blanchet
parents: 49665
diff changeset
    93
    Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct
blanchet
parents: 49665
diff changeset
    94
  | _ => Conv.concl_conv ~1 cv ct);
blanchet
parents: 49665
diff changeset
    95
54922
blanchet
parents: 54837
diff changeset
    96
fun co_induct_inst_as_projs ctxt k thm =
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    97
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    98
    val fs = Term.add_vars (Thm.prop_of thm) []
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
    99
      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   100
    fun mk_inst (xi, T) = (xi, Thm.cterm_of ctxt (mk_proj T (num_binder_types T) k));
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   101
  in
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   102
    infer_instantiate ctxt (map mk_inst fs) thm
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   103
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   104
54922
blanchet
parents: 54837
diff changeset
   105
val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   106
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   107
fun mk_case_transfer_tac ctxt rel_case cases =
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   108
  let val n = length (tl (Thm.prems_of rel_case)) in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   109
    REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   110
    HEADGOAL (etac ctxt rel_case) THEN
58093
6f37a300c82b generate 'case_transfer' for (co)datatypes
desharna
parents: 58044
diff changeset
   111
    ALLGOALS (hyp_subst_tac ctxt) THEN
6f37a300c82b generate 'case_transfer' for (co)datatypes
desharna
parents: 58044
diff changeset
   112
    unfold_thms_tac ctxt cases THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   113
    ALLGOALS (fn k => (select_prem_tac ctxt n (dtac ctxt asm_rl) k) k) THEN
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   114
    ALLGOALS (REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt rel_funD THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   115
      (assume_tac ctxt THEN' etac ctxt thin_rl ORELSE' rtac ctxt refl)) THEN' assume_tac ctxt)
58093
6f37a300c82b generate 'case_transfer' for (co)datatypes
desharna
parents: 58044
diff changeset
   116
  end;
6f37a300c82b generate 'case_transfer' for (co)datatypes
desharna
parents: 58044
diff changeset
   117
58327
a147bd03cad0 make 'ctr_transfer' tactic more robust
desharna
parents: 58326
diff changeset
   118
fun mk_ctr_transfer_tac ctxt rel_intros rel_eqs =
58095
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   119
  HEADGOAL Goal.conjunction_tac THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   120
  ALLGOALS (REPEAT o (resolve_tac ctxt (rel_funI :: rel_intros) THEN'
62535
cb262f03ac12 strengthened tactic
blanchet
parents: 62335
diff changeset
   121
    TRY o (REPEAT_DETERM1 o (SELECT_GOAL (unfold_thms_tac ctxt rel_eqs) THEN'
cb262f03ac12 strengthened tactic
blanchet
parents: 62335
diff changeset
   122
      (assume_tac ctxt ORELSE' hyp_subst_tac ctxt THEN' rtac ctxt refl)))));
58095
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   123
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   124
fun mk_disc_transfer_tac ctxt rel_sel exhaust_disc distinct_disc =
58095
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   125
  let
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   126
    fun last_disc_tac iffD =
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   127
      HEADGOAL (rtac ctxt (rotate_prems ~1 exhaust_disc) THEN' assume_tac ctxt THEN'
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   128
      REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt (rotate_prems 1 iffD) THEN'
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   129
        assume_tac ctxt THEN' rotate_tac ~1 THEN'
60757
c09598a97436 prefer tactics with explicit context;
wenzelm
parents: 60752
diff changeset
   130
        etac ctxt (rotate_prems 1 notE) THEN' eresolve_tac ctxt distinct_disc));
58095
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   131
  in
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   132
    HEADGOAL Goal.conjunction_tac THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   133
    REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI THEN' dtac ctxt (rel_sel RS iffD1) THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   134
      REPEAT_DETERM o (etac ctxt conjE) THEN' (assume_tac ctxt ORELSE' rtac ctxt iffI))) THEN
58095
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   135
    TRY (last_disc_tac iffD2) THEN TRY (last_disc_tac iffD1)
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   136
  end;
b280d4028443 generate 'disc_transfer' for (co)datatypes
desharna
parents: 58093
diff changeset
   137
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
   138
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   139
  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac ctxt sumEN') THEN
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   140
  HEADGOAL (EVERY' (maps (fn k => [select_prem_tac ctxt n (rotate_tac 1) k,
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   141
    REPEAT_DETERM o dtac ctxt meta_spec, etac ctxt meta_mp, assume_tac ctxt]) (1 upto n)));
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
   142
49501
acc9635a644a renamed "fld"/"unf" to "ctor"/"dtor"
blanchet
parents: 49484
diff changeset
   143
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   144
  HEADGOAL (rtac ctxt iffI THEN'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   145
    EVERY' (@{map 3} (fn cTs => fn cx => fn th =>
60801
7664e0916eec tuned signature;
wenzelm
parents: 60784
diff changeset
   146
      dtac ctxt (Thm.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
52324
blanchet
parents: 52323
diff changeset
   147
      SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   148
      assume_tac ctxt) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   149
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   150
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
   151
  unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   152
  HEADGOAL (rtac ctxt @{thm sum.distinct(1)});
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
   153
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   154
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
   155
  unfold_thms_tac ctxt [ctr_def] THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   156
  HEADGOAL (rtac ctxt (ctor_inject RS ssubst)) THEN
56765
644f0d4820a1 cleaner 'rel_inject' theorems
blanchet
parents: 56245
diff changeset
   157
  unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject prod.inject conj_assoc}) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   158
  HEADGOAL (rtac ctxt refl);
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   159
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   160
val rec_unfold_thms =
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
   161
  @{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
   162
      case_unit_Unity} @ sumprod_thms_map;
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   163
58734
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   164
fun mk_co_rec_o_map_tac ctxt co_rec_def pre_map_defs map_ident0s abs_inverses xtor_co_rec_o_map =
58732
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   165
  let
58734
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   166
    val rec_o_map_simps = @{thms o_def[abs_def] id_def case_prod_app case_sum_map_sum map_sum.simps
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   167
      case_prod_map_prod id_bnf_def map_prod_simp map_sum_if_distrib_then map_sum_if_distrib_else
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   168
      if_distrib[THEN sym]};
58732
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   169
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   170
    HEADGOAL (subst_tac ctxt (SOME [1, 2]) [co_rec_def] THEN'
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   171
      rtac ctxt (xtor_co_rec_o_map RS trans) THEN'
58732
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   172
      CONVERSION Thm.eta_long_conversion THEN'
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   173
      asm_simp_tac (ss_only (pre_map_defs @ distinct Thm.eq_thm_prop (map_ident0s @ abs_inverses) @
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   174
        rec_o_map_simps) ctxt))
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   175
  end;
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58676
diff changeset
   176
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   177
fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec fp_abs_inverse abs_inverse ctr_def ctxt =
58359
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   178
  HEADGOAL ((if is_def_looping ctr_def then subst_tac ctxt NONE
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   179
    else SELECT_GOAL o unfold_thms_tac ctxt) [ctr_def]) THEN
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58327
diff changeset
   180
  unfold_thms_tac ctxt (ctor_rec :: fp_abs_inverse :: abs_inverse :: rec_defs @
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   181
    pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac ctxt refl);
49205
674f04c737e0 implemented "mk_iter_or_rec_tac"
blanchet
parents: 49161
diff changeset
   182
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   183
fun mk_rec_transfer_tac ctxt nn ns actives passives xssss rec_defs ctor_rec_transfers rel_pre_T_defs
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   184
    rel_eqs =
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   185
  let
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   186
    val ctor_rec_transfers' =
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   187
      map (infer_instantiate' ctxt (map SOME (passives @ actives))) ctor_rec_transfers;
58507
blanchet
parents: 58455
diff changeset
   188
    val total_n = Integer.sum ns;
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   189
    val True = @{term True};
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   190
  in
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   191
    HEADGOAL Goal.conjunction_tac THEN
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   192
    EVERY (map (fn ctor_rec_transfer =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   193
        REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   194
        unfold_thms_tac ctxt rec_defs THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   195
        HEADGOAL (etac ctxt (mk_rel_funDN_rotated (nn + 1) ctor_rec_transfer)) THEN
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   196
        unfold_thms_tac ctxt rel_pre_T_defs THEN
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   197
        EVERY (fst (@{fold_map 2} (fn k => fn xsss => fn acc =>
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   198
            rpair (k + acc)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   199
            (HEADGOAL (rtac ctxt (mk_rel_funDN_rotated 2 @{thm comp_transfer})) THEN
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   200
             HEADGOAL (rtac ctxt @{thm vimage2p_rel_fun}) THEN
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   201
             unfold_thms_tac ctxt rel_eqs THEN
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   202
             EVERY (@{map 2} (fn n => fn xss =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   203
                 REPEAT_DETERM (HEADGOAL (resolve_tac ctxt
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   204
                   [mk_rel_funDN 2 case_sum_transfer_eq, mk_rel_funDN 2 case_sum_transfer])) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   205
                 HEADGOAL (select_prem_tac ctxt total_n (dtac ctxt asm_rl) (acc + n)) THEN
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   206
                 HEADGOAL (SELECT_GOAL (HEADGOAL
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   207
                   (REPEAT_DETERM o (assume_tac ctxt ORELSE' resolve_tac ctxt
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   208
                       [mk_rel_funDN 1 case_prod_transfer_eq,
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   209
                        mk_rel_funDN 1 case_prod_transfer,
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   210
                        rel_funI]) THEN_ALL_NEW
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   211
                    (Subgoal.FOCUS (fn {prems, ...} =>
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   212
                       let val thm = prems
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   213
                         |> permute_like (op =) (True :: flat xss) (True :: flat_rec_arg_args xss)
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   214
                         |> Library.foldl1 (fn (acc, elem) => elem RS (acc RS rel_funD))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   215
                       in HEADGOAL (rtac ctxt thm) end) ctxt)))))
58966
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   216
               (1 upto k) xsss)))
0297e1277ed2 make 'rec_transfer' tactic more robust
desharna
parents: 58734
diff changeset
   217
          ns xssss 0)))
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   218
      ctor_rec_transfers')
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   219
  end;
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58417
diff changeset
   220
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55944
diff changeset
   221
val corec_unfold_thms = @{thms id_def} @ sumprod_thms_map;
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49680
diff changeset
   222
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   223
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
   224
  let
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   225
    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
   226
      @{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
   227
  in
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   228
    unfold_thms_tac ctxt (ctr_def :: corec_defs) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   229
    HEADGOAL (rtac ctxt (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   230
    HEADGOAL (rtac ctxt refl ORELSE' rtac ctxt (@{thm unit_eq} RS arg_cong))
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   231
  end;
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49211
diff changeset
   232
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   233
fun mk_corec_disc_iff_tac case_splits' corecs discs ctxt =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   234
  EVERY (@{map 3} (fn case_split_tac => fn corec_thm => fn disc =>
55867
79b915f26533 rationalized internals
blanchet
parents: 55803
diff changeset
   235
      HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN
52324
blanchet
parents: 52323
diff changeset
   236
      HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   237
      (if is_refl disc then all_tac else HEADGOAL (rtac ctxt disc)))
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   238
    (map (rtac ctxt) case_splits' @ [K all_tac]) corecs discs);
49482
e6d6869eed08 generate coiter_iff and corec_iff theorems
blanchet
parents: 49463
diff changeset
   239
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   240
fun mk_corec_transfer_tac ctxt actives passives type_definitions corec_defs dtor_corec_transfers
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   241
    rel_pre_T_defs rel_eqs pgs pss qssss gssss =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   242
  let
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   243
    val num_pgs = length pgs;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   244
    fun prem_no_of x = 1 + find_index (curry (op =) x) pgs;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   245
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   246
    val Inl_Inr_Pair_tac = REPEAT_DETERM o (resolve_tac ctxt
58968
d09bbd2642e2 make 'corec_transfer' tactic more robust
desharna
parents: 58966
diff changeset
   247
      [mk_rel_funDN 1 @{thm Inl_transfer},
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63068
diff changeset
   248
       mk_rel_funDN 1 @{thm Inl_transfer[of "op =" "op =", simplified sum.rel_eq]},
58968
d09bbd2642e2 make 'corec_transfer' tactic more robust
desharna
parents: 58966
diff changeset
   249
       mk_rel_funDN 1 @{thm Inr_transfer},
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63068
diff changeset
   250
       mk_rel_funDN 1 @{thm Inr_transfer[of "op =" "op =", simplified sum.rel_eq]},
58968
d09bbd2642e2 make 'corec_transfer' tactic more robust
desharna
parents: 58966
diff changeset
   251
       mk_rel_funDN 2 @{thm Pair_transfer},
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63068
diff changeset
   252
       mk_rel_funDN 2 @{thm Pair_transfer[of "op =" "op =", simplified prod.rel_eq]}]);
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   253
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   254
    fun mk_unfold_If_tac total pos =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   255
      HEADGOAL (Inl_Inr_Pair_tac THEN'
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   256
        rtac ctxt (mk_rel_funDN 3 @{thm If_transfer}) THEN'
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   257
        select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   258
        dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   259
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   260
    fun mk_unfold_Inl_Inr_Pair_tac total pos =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   261
      HEADGOAL (Inl_Inr_Pair_tac THEN'
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   262
        select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   263
        dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   264
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   265
    fun mk_unfold_arg_tac qs gs =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   266
      EVERY (map (mk_unfold_If_tac num_pgs o prem_no_of) qs) THEN
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   267
      EVERY (map (mk_unfold_Inl_Inr_Pair_tac num_pgs o prem_no_of) gs);
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   268
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   269
    fun mk_unfold_ctr_tac type_definition qss gss =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   270
      HEADGOAL (rtac ctxt (mk_rel_funDN 1 (@{thm Abs_transfer} OF
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   271
        [type_definition, type_definition])) THEN' Inl_Inr_Pair_tac) THEN
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   272
      (case (qss, gss) of
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   273
        ([], []) => HEADGOAL (rtac ctxt refl)
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   274
      | _ => EVERY (map2 mk_unfold_arg_tac qss gss));
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   275
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   276
    fun mk_unfold_type_tac type_definition ps qsss gsss =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   277
      let
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   278
        val p_tacs = map (mk_unfold_If_tac num_pgs o prem_no_of) ps;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   279
        val qg_tacs = map2 (mk_unfold_ctr_tac type_definition) qsss gsss;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   280
        fun mk_unfold_ty [] [qg_tac] = qg_tac
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   281
          | mk_unfold_ty (p_tac :: p_tacs) (qg_tac :: qg_tacs) =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   282
            p_tac THEN qg_tac THEN mk_unfold_ty p_tacs qg_tacs
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   283
      in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   284
        HEADGOAL (rtac ctxt rel_funI) THEN mk_unfold_ty p_tacs qg_tacs
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   285
      end;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   286
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   287
    fun mk_unfold_corec_type_tac dtor_corec_transfer corec_def =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   288
      let
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   289
        val active :: actives' = actives;
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   290
        val dtor_corec_transfer' =
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   291
          infer_instantiate' ctxt
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   292
            (SOME active :: map SOME passives @ map SOME actives') dtor_corec_transfer;
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   293
      in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   294
        HEADGOAL Goal.conjunction_tac THEN REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   295
        unfold_thms_tac ctxt [corec_def] THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   296
        HEADGOAL (etac ctxt (mk_rel_funDN_rotated (1 + length actives) dtor_corec_transfer')) THEN
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   297
        unfold_thms_tac ctxt (rel_pre_T_defs @ rel_eqs)
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   298
      end;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   299
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   300
    fun mk_unfold_prop_tac dtor_corec_transfer corec_def =
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   301
      mk_unfold_corec_type_tac dtor_corec_transfer corec_def THEN
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   302
      EVERY (@{map 4} mk_unfold_type_tac type_definitions pss qssss gssss);
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   303
  in
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   304
    HEADGOAL Goal.conjunction_tac THEN
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   305
    EVERY (map2 mk_unfold_prop_tac dtor_corec_transfers corec_defs)
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   306
  end;
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   307
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   308
fun solve_prem_prem_tac ctxt =
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   309
  REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE'
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   310
    rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE' hyp_subst_tac ctxt ORELSE'
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   311
    resolve_tac ctxt @{thms disjI1 disjI2}) THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   312
  (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' rtac ctxt @{thm singletonI});
49426
6d05b8452cf3 got rid of one "auto" in induction tactic
blanchet
parents: 49391
diff changeset
   313
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   314
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
   315
    pre_set_defs =
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   316
  HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk,
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   317
    etac ctxt meta_mp,
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   318
    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
   319
      sumprod_thms_set)),
52324
blanchet
parents: 52323
diff changeset
   320
    solve_prem_prem_tac ctxt]) (rev kks)));
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   321
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   322
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
   323
    kks =
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   324
  let val r = length kks in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   325
    HEADGOAL (EVERY' [select_prem_tac ctxt n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   326
      REPEAT_DETERM_N m o (dtac ctxt meta_spec THEN' rotate_tac ~1)]) THEN
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   327
    EVERY [REPEAT_DETERM_N r
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   328
        (HEADGOAL (rotate_tac ~1 THEN' dtac ctxt meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   329
      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL (assume_tac ctxt),
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   330
      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
   331
        pre_set_defs]
49391
3a96797fd53e tuned induction tactic
blanchet
parents: 49389
diff changeset
   332
  end;
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   333
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   334
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
   335
    pre_set_defss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   336
  let val n = Integer.sum ns in
58359
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   337
    (if exists is_def_looping ctr_defs then
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   338
       EVERY (map (fn def => HEADGOAL (subst_asm_tac ctxt NONE [def])) ctr_defs)
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   339
     else
3782c7b0d1cc avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
blanchet
parents: 58353
diff changeset
   340
       unfold_thms_tac ctxt ctr_defs) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   341
    HEADGOAL (rtac ctxt ctor_induct') THEN co_induct_inst_as_projs_tac ctxt 0 THEN
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   342
    EVERY (@{map 4} (EVERY oooo @{map 3} o
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   343
        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
   344
      pre_set_defss mss (unflat mss (1 upto n)) kkss)
49368
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   345
  end;
df359ce891ac added induct tactic
blanchet
parents: 49363
diff changeset
   346
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   347
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
   348
    discs sels =
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51766
diff changeset
   349
  hyp_subst_tac ctxt THEN'
49665
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   350
  CONVERSION (hhf_concl_conv
869c7a2e2945 tuned tactic
traytel
parents: 49643
diff changeset
   351
    (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
   352
  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
   353
  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
   354
    sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   355
  (assume_tac ctxt ORELSE' REPEAT o etac ctxt conjE THEN'
56765
644f0d4820a1 cleaner 'rel_inject' theorems
blanchet
parents: 56245
diff changeset
   356
     full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   357
     REPEAT o etac ctxt conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   358
     REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' assume_tac ctxt));
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   359
52966
blanchet
parents: 52659
diff changeset
   360
fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
54198
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   361
  let
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   362
    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
   363
      |> distinct Thm.eq_thm_prop;
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   364
  in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   365
    hyp_subst_tac ctxt THEN' REPEAT o etac ctxt conjE THEN'
54198
4fadf746f2d5 got rid of annoying duplicate rewrite rule warnings
blanchet
parents: 53690
diff changeset
   366
    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
   367
  end;
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   368
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   369
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
   370
    dtor_ctor exhaust ctr_defs discss selss =
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   371
  let val ks = 1 upto n in
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   372
    EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   373
        select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, dtac ctxt meta_spec, dtac ctxt meta_mp,
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   374
        assume_tac ctxt, rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   375
        hyp_subst_tac ctxt] @
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   376
      @{map 4} (fn k => fn ctr_def => fn discs => fn sels =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   377
        EVERY' ([rtac ctxt 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
   378
          map2 (fn k' => fn discs' =>
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   379
            if k' = k then
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   380
              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
   381
                dtor_ctor ctr_def discs sels
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
   382
            else
52966
blanchet
parents: 52659
diff changeset
   383
              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
   384
  end;
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   385
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   386
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
   387
    dtor_ctors exhausts ctr_defss discsss selsss =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   388
  HEADGOAL (rtac ctxt dtor_coinduct' THEN'
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   389
    EVERY' (@{map 10} (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55642
diff changeset
   390
      (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
   391
      selsss));
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49589
diff changeset
   392
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   393
fun mk_map_tac ctxt abs_inverses pre_map_def dtor_ctor fp_map' live_nesting_map_ident0s ctr_defs' =
63840
eb6d2aace13a derive maps forward
blanchet
parents: 63835
diff changeset
   394
  TRYALL Goal.conjunction_tac THEN
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   395
  unfold_thms_tac ctxt (pre_map_def :: dtor_ctor :: fp_map' :: abs_inverses @
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   396
    live_nesting_map_ident0s @
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   397
    ctr_defs' @ sumprod_thms_map @ @{thms o_def[abs_def] id_def}) THEN
63840
eb6d2aace13a derive maps forward
blanchet
parents: 63835
diff changeset
   398
  ALLGOALS (rtac ctxt refl);
eb6d2aace13a derive maps forward
blanchet
parents: 63835
diff changeset
   399
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   400
fun mk_map_disc_iff_tac ctxt ct exhaust discs maps =
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   401
  TRYALL Goal.conjunction_tac THEN
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   402
  ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   403
    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   404
  unfold_thms_tac ctxt maps THEN
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   405
  unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   406
    handle THM _ => thm RS eqTrueI) discs) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   407
  ALLGOALS (rtac ctxt refl ORELSE' rtac ctxt TrueI);
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   408
58326
7e142efcee1a make 'rel_sel' and 'map_sel' tactics more robust
desharna
parents: 58181
diff changeset
   409
fun mk_map_sel_tac ctxt ct exhaust discs maps sels map_id0s =
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   410
  TRYALL Goal.conjunction_tac THEN
63841
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   411
  ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   412
    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   413
  unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   414
    @{thms not_True_eq_False not_False_eq_True}) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   415
  TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   416
  unfold_thms_tac ctxt (@{thm id_apply} :: maps @ sels @ map_id0s) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   417
  ALLGOALS (rtac ctxt refl);
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   418
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   419
fun mk_rel_tac ctxt abs_inverses pre_rel_def dtor_ctor fp_rel live_nesting_rel_eqs ctr_defs' =
63841
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   420
  TRYALL Goal.conjunction_tac THEN
63842
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   421
  unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: fp_rel :: abs_inverses @ live_nesting_rel_eqs @
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   422
    ctr_defs' @ sumprod_thms_rel @ @{thms vimage2p_def o_apply sum.inject
f738df816abf strengthened tactics
blanchet
parents: 63841
diff changeset
   423
      sum.distinct(1)[THEN eq_False[THEN iffD2]] not_False_eq_True}) THEN
63841
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   424
  ALLGOALS (resolve_tac ctxt [TrueI, refl]);
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   425
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   426
fun mk_rel_case_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts rel_eqs =
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   427
  HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct1] exhaust) THEN_ALL_NEW
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   428
    rtac ctxt (infer_instantiate' ctxt [SOME ct2] exhaust) THEN_ALL_NEW
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   429
      hyp_subst_tac ctxt) THEN
61344
ebf296fe88d7 generate 'case_transfer' unconditionally
blanchet
parents: 60801
diff changeset
   430
  unfold_thms_tac ctxt (rel_eqs @ injects @ rel_injects @
ebf296fe88d7 generate 'case_transfer' unconditionally
blanchet
parents: 60801
diff changeset
   431
    @{thms conj_imp_eq_imp_imp simp_thms(6) True_implies_equals} @
57529
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   432
    map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @
5e83df79eaa0 refactor some tactics
desharna
parents: 57528
diff changeset
   433
    map (fn thm => thm RS eqTrueI) rel_injects) THEN
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   434
  TRYALL (assume_tac ctxt ORELSE' etac ctxt FalseE ORELSE'
61344
ebf296fe88d7 generate 'case_transfer' unconditionally
blanchet
parents: 60801
diff changeset
   435
    (REPEAT_DETERM o dtac ctxt meta_spec THEN'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   436
     TRY o filter_prems_tac ctxt
57562
c1238062184b fix rel_cases
desharna
parents: 57558
diff changeset
   437
       (forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN'
61344
ebf296fe88d7 generate 'case_transfer' unconditionally
blanchet
parents: 60801
diff changeset
   438
     REPEAT_DETERM o (dtac ctxt meta_mp THEN' rtac ctxt refl) THEN'
ebf296fe88d7 generate 'case_transfer' unconditionally
blanchet
parents: 60801
diff changeset
   439
     (assume_tac ctxt ORELSE' Goal.assume_rule_tac ctxt)));
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57471
diff changeset
   440
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   441
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
   442
    dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses nesting_rel_eqs =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   443
  rtac ctxt dtor_rel_coinduct 1 THEN
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58507
diff changeset
   444
   EVERY (@{map 11} (fn ct => fn assm => fn exhaust => fn discs => fn sels => fn ctor_defs =>
57670
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   445
     fn dtor_ctor => fn ctor_inject => fn abs_inject => fn rel_pre_def => fn abs_inverse =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   446
      (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   447
         (dtac ctxt (rotate_prems ~1 (infer_instantiate' ctxt [NONE, NONE, NONE, NONE, SOME ct]
57670
d7b15b99f93c compile
blanchet
parents: 57668
diff changeset
   448
            @{thm arg_cong2} RS iffD1)) THEN'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   449
          assume_tac ctxt THEN' assume_tac ctxt THEN' hyp_subst_tac ctxt THEN' dtac ctxt assm THEN'
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   450
          REPEAT_DETERM o etac ctxt conjE))) 1 THEN
57668
blanchet
parents: 57563
diff changeset
   451
      unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN
57528
9afc832252a3 refactor some tactics
desharna
parents: 57525
diff changeset
   452
      unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject ::
58128
43a1ba26a8cb renamed BNF theories
blanchet
parents: 58095
diff changeset
   453
        abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   454
        @{thms id_bnf_def rel_sum_simps rel_prod_inject vimage2p_def Inl_Inr_False
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   455
          iffD2[OF eq_False Inr_not_Inl] sum.inject prod.inject}) THEN
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   456
      REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN'
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   457
        (REPEAT_DETERM o rtac ctxt conjI) THEN' (rtac ctxt refl ORELSE' assume_tac ctxt))))
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   458
    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
   459
      abs_inverses);
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57152
diff changeset
   460
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   461
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
   462
    rel_pre_list_defs Abs_inverses nesting_rel_eqs =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   463
  rtac ctxt ctor_rel_induct 1 THEN EVERY (@{map 6} (fn cterm => fn exhaust => fn ctor_defs =>
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   464
      fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   465
        HEADGOAL (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   466
          (rtac ctxt (infer_instantiate' ctxt (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2}
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   467
              RS iffD2)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   468
            THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' TRY o resolve_tac ctxt assms))) THEN
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   469
        unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   470
          @{thms id_bnf_def vimage2p_def}) THEN
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   471
        TRYALL (hyp_subst_tac ctxt) THEN
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 61760
diff changeset
   472
        unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_inject Inl_Inr_False
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   473
          Inr_Inl_False  sum.inject prod.inject}) THEN
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   474
        TRYALL (rtac ctxt refl ORELSE' etac ctxt FalseE ORELSE'
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   475
          (REPEAT_DETERM o etac ctxt conjE) THEN' assume_tac ctxt))
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   476
    cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses);
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57399
diff changeset
   477
58326
7e142efcee1a make 'rel_sel' and 'map_sel' tactics more robust
desharna
parents: 58181
diff changeset
   478
fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts rel_eqs =
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   479
  HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct1] exhaust) THEN_ALL_NEW
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   480
    rtac ctxt (infer_instantiate' ctxt [SOME ct2] exhaust) THEN_ALL_NEW
57563
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   481
      hyp_subst_tac ctxt) THEN
59270
blanchet
parents: 59269
diff changeset
   482
  unfold_thms_tac ctxt (sels @ rel_injects @ rel_eqs @
blanchet
parents: 59269
diff changeset
   483
    @{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @ ((discs @ distincts) RL [eqTrueI, eqFalseI]) @
blanchet
parents: 59269
diff changeset
   484
    (rel_injects RL [eqTrueI]) @ (rel_distincts RL [eqFalseI])) THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   485
  TRYALL (resolve_tac ctxt [TrueI, refl]);
57563
e3e7c86168b4 generate 'rel_sel' theorem for (co)datatypes
desharna
parents: 57562
diff changeset
   486
58676
cdf84b6e1297 generate 'sel_transfer' for (co)datatypes
desharna
parents: 58634
diff changeset
   487
fun mk_sel_transfer_tac ctxt n sel_defs case_transfer =
cdf84b6e1297 generate 'sel_transfer' for (co)datatypes
desharna
parents: 58634
diff changeset
   488
  TRYALL Goal.conjunction_tac THEN
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 62535
diff changeset
   489
  unfold_thms_tac ctxt (map (Local_Defs.abs_def_rule ctxt) sel_defs) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   490
  ALLGOALS (rtac ctxt (mk_rel_funDN n case_transfer) THEN_ALL_NEW
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   491
    REPEAT_DETERM o (assume_tac ctxt ORELSE' rtac ctxt rel_funI));
58676
cdf84b6e1297 generate 'sel_transfer' for (co)datatypes
desharna
parents: 58634
diff changeset
   492
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57893
diff changeset
   493
fun mk_set_sel_tac ctxt ct exhaust discs sels sets =
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   494
  TRYALL Goal.conjunction_tac THEN
63841
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   495
  ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   496
    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   497
  unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   498
    @{thms not_True_eq_False not_False_eq_True}) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   499
  TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   500
  unfold_thms_tac ctxt (sels @ sets) THEN
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   501
  ALLGOALS (REPEAT o (resolve_tac ctxt @{thms UnI1 UnI2 imageI} ORELSE'
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   502
      eresolve_tac ctxt @{thms UN_I UN_I[rotated] imageE} ORELSE'
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   503
      hyp_subst_tac ctxt) THEN'
813a574da746 derive relator properties forward
blanchet
parents: 63840
diff changeset
   504
    (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
57152
de1ed2c1c3bf generate 'sel_set' theorem for (co)datatypes
desharna
parents: 57046
diff changeset
   505
57893
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
   506
fun mk_set_cases_tac ctxt ct assms exhaust sets =
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   507
  HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust)
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   508
    THEN_ALL_NEW hyp_subst_tac ctxt) THEN
57893
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
   509
  unfold_thms_tac ctxt sets THEN
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
   510
  REPEAT_DETERM (HEADGOAL
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   511
    (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
57893
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
   512
     hyp_subst_tac ctxt ORELSE'
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   513
     SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   514
       assume_tac ctxt)))));
57893
7bc128647d6e generate 'set_cases' theorem for (co)datatypes
desharna
parents: 57891
diff changeset
   515
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
   516
fun mk_set_intros_tac ctxt sets =
59856
ed0ca9029021 export more low-level theorems in data structure (partly for 'corec')
blanchet
parents: 59794
diff changeset
   517
  TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt sets THEN
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
   518
  TRYALL (REPEAT o
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   519
    (resolve_tac ctxt @{thms UnI1 UnI2} ORELSE'
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   520
     eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN'
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   521
     (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
57891
d23a85b59dec generate 'set_intros' theorem for (co)datatypes
desharna
parents: 57882
diff changeset
   522
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   523
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
   524
    Abs_pre_inverses =
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   525
  let
58417
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58359
diff changeset
   526
    val assms_tac =
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58359
diff changeset
   527
      let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58359
diff changeset
   528
        fold (curry (op ORELSE')) (map (fn thm =>
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   529
            funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' assume_tac ctxt)
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   530
              (rtac ctxt thm)) assms')
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59856
diff changeset
   531
          (etac ctxt FalseE)
58417
fa50722ad6cb make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
desharna
parents: 58359
diff changeset
   532
      end;
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   533
    val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   534
      |> map2 (fn ct => infer_instantiate' ctxt [NONE, SOME ct]) cts;
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   535
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   536
    ALLGOALS (resolve_tac ctxt dtor_set_inducts) THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   537
    TRYALL (resolve_tac ctxt exhausts' THEN_ALL_NEW
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59270
diff changeset
   538
      (resolve_tac ctxt (map (fn ct => refl RS
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60757
diff changeset
   539
         infer_instantiate' ctxt (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   540
        THEN' assume_tac ctxt THEN' hyp_subst_tac ctxt)) THEN
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   541
    unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   542
      @{thms id_bnf_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert Un_empty_left
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   543
        Un_empty_right empty_iff singleton_iff}) THEN
61760
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   544
    REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE'
1647bb489522 tuned whitespace
blanchet
parents: 61344
diff changeset
   545
      eresolve_tac ctxt @{thms UN_E UnE singletonE} ORELSE' assms_tac))
57700
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   546
  end;
a2c4adb839a9 generate 'set_induct' theorem for codatatypes
desharna
parents: 57670
diff changeset
   547
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff changeset
   548
end;