src/HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
author blanchet
Thu, 02 Jan 2014 20:25:40 +0100
changeset 54916 aa891e065af1
parent 54915 61284fe0536a
child 54917 a426e38a8a7e
permissions -rw-r--r--
simplified code
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54174
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     3
    Copyright   2013
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     4
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54174
diff changeset
     5
Tactics for corecursor sugar.
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     6
*)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     7
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54174
diff changeset
     8
signature BNF_GFP_REC_SUGAR_TACTICS =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     9
sig
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
    10
  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
    11
  val mk_primcorec_code_of_raw_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
    12
  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
53693
71b020d161c5 new tactics for constructor view
blanchet
parents: 53692
diff changeset
    13
  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
71b020d161c5 new tactics for constructor view
blanchet
parents: 53692
diff changeset
    14
    tactic
54910
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    15
  val mk_primcorec_disc_iff_tac: Proof.context -> thm -> thm -> thm list -> thm list -> tactic
54844
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    16
  val mk_primcorec_exhaust_tac: int -> thm -> tactic
54043
58a0f8726558 renamings
blanchet
parents: 54042
diff changeset
    17
  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
54613
985f8b49c050 more work towards "exhaustive"
panny
parents: 54279
diff changeset
    18
    thm list -> int list -> thm list -> thm option -> tactic
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
    19
  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
54279
3ffb74b52ed6 removed dead code
blanchet
parents: 54246
diff changeset
    20
    thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    21
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    22
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54174
diff changeset
    23
structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    24
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    25
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    26
open BNF_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    27
open BNF_Tactics
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    28
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
    29
val atomize_conjL = @{thm atomize_conjL};
53905
blanchet
parents: 53904
diff changeset
    30
val falseEs = @{thms not_TrueE FalseE};
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
    31
val Let_def = @{thm Let_def};
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
    32
val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
53902
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    33
val split_if = @{thm split_if};
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    34
val split_if_asm = @{thm split_if_asm};
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    35
val split_connectI = @{thms allI impI conjI};
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    36
54844
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    37
fun mk_primcorec_exhaust_tac n nchotomy =
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    38
  let val ks = 1 upto n in
54905
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    39
    HEADGOAL (atac ORELSE'
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    40
      cut_tac nchotomy THEN'
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    41
      EVERY' (map (fn k =>
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    42
          (if k < n then etac disjE else K all_tac) THEN'
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    43
          REPEAT o (etac conjE THEN' dtac meta_mp THEN' atac) THEN'
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    44
          dtac meta_mp THEN' atac THEN' atac)
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    45
        ks))
54844
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    46
  end;
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    47
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
    48
fun mk_primcorec_assumption_tac ctxt discIs =
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
    49
  SELECT_GOAL (unfold_thms_tac ctxt
54069
3fd3b1683d2b strengthen the tactics to bring them in sync with the simplifications taking place in the code (and weaken the simplifications a bit -- let's not deal with implies)
blanchet
parents: 54044
diff changeset
    50
      @{thms not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
53926
9fc9a59ad579 strengthened tactic
blanchet
parents: 53922
diff changeset
    51
    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
54117
32730ba3ab85 strengthened tactic to deal with 'False ==> ...'
blanchet
parents: 54103
diff changeset
    52
    eresolve_tac falseEs ORELSE'
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
    53
    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
53929
8c5aaf557421 further strengthening of tactics
blanchet
parents: 53926
diff changeset
    54
    dresolve_tac discIs THEN' atac ORELSE'
8c5aaf557421 further strengthening of tactics
blanchet
parents: 53926
diff changeset
    55
    etac notE THEN' atac ORELSE'
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
    56
    etac disjE))));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    57
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    58
fun mk_primcorec_same_case_tac m =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    59
  HEADGOAL (if m = 0 then rtac TrueI
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    60
    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    61
54069
3fd3b1683d2b strengthen the tactics to bring them in sync with the simplifications taking place in the code (and weaken the simplifications a bit -- let's not deal with implies)
blanchet
parents: 54044
diff changeset
    62
fun mk_primcorec_different_case_tac ctxt m excl =
3fd3b1683d2b strengthen the tactics to bring them in sync with the simplifications taking place in the code (and weaken the simplifications a bit -- let's not deal with implies)
blanchet
parents: 54044
diff changeset
    63
  HEADGOAL (if m = 0 then mk_primcorec_assumption_tac ctxt []
3fd3b1683d2b strengthen the tactics to bring them in sync with the simplifications taking place in the code (and weaken the simplifications a bit -- let's not deal with implies)
blanchet
parents: 54044
diff changeset
    64
    else dtac excl THEN' (REPEAT_DETERM_N (m - 1) o atac) THEN' mk_primcorec_assumption_tac ctxt []);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    65
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    66
fun mk_primcorec_cases_tac ctxt k m exclsss =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    67
  let val n = length exclsss in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    68
    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
54069
3fd3b1683d2b strengthen the tactics to bring them in sync with the simplifications taking place in the code (and weaken the simplifications a bit -- let's not deal with implies)
blanchet
parents: 54044
diff changeset
    69
        | [excl] => mk_primcorec_different_case_tac ctxt m excl)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    70
      (take k (nth exclsss (k - 1))))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    71
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    72
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    73
fun mk_primcorec_prelude ctxt defs thm =
54024
07ab4fd922c2 strengthen tactic w.r.t. let
blanchet
parents: 54018
diff changeset
    74
  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
07ab4fd922c2 strengthen tactic w.r.t. let
blanchet
parents: 54018
diff changeset
    75
  unfold_thms_tac ctxt @{thms Let_def split};
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    76
53706
9e28c41e3595 more primcorec tactics
blanchet
parents: 53702
diff changeset
    77
fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
9e28c41e3595 more primcorec tactics
blanchet
parents: 53702
diff changeset
    78
  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    79
54910
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    80
fun mk_primcorec_disc_iff_tac ctxt fun_exhaust fun_disc fun_discs disc_excludes =
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    81
  HEADGOAL (rtac iffI THEN'
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    82
    rtac fun_exhaust THEN'
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    83
    EVERY' (map (fn fun_disc' =>
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    84
        if Thm.eq_thm (fun_disc', fun_disc) then
54916
aa891e065af1 simplified code
blanchet
parents: 54915
diff changeset
    85
          REPEAT_DETERM o etac conjI THEN' (atac ORELSE' rtac TrueI)
54910
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    86
        else
54915
61284fe0536a prevent tactic from getting out of sync and wrongly attack next subgoal
blanchet
parents: 54910
diff changeset
    87
          dtac fun_disc' THEN' (REPEAT_DETERM o atac) THEN'
61284fe0536a prevent tactic from getting out of sync and wrongly attack next subgoal
blanchet
parents: 54910
diff changeset
    88
          DETERM o (dresolve_tac disc_excludes THEN' etac notE THEN' atac))
54910
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    89
      fun_discs) THEN'
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
    90
    rtac (unfold_thms ctxt [atomize_conjL] fun_disc) THEN_MAYBE' atac);
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
    91
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
    92
fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms map_idents map_comps fun_sel k m
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
    93
    exclsss =
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
    94
  mk_primcorec_prelude ctxt defs (fun_sel RS trans) THEN
53693
71b020d161c5 new tactics for constructor view
blanchet
parents: 53692
diff changeset
    95
  mk_primcorec_cases_tac ctxt k m exclsss THEN
54018
bd2e127389f2 strengthened tactic for right-hand sides involving lambdas
blanchet
parents: 53961
diff changeset
    96
  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
53905
blanchet
parents: 53904
diff changeset
    97
    eresolve_tac falseEs ORELSE'
53902
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    98
    resolve_tac split_connectI ORELSE'
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    99
    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
   100
    Splitter.split_tac (split_if :: splits) ORELSE'
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
   101
    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
   102
    etac notE THEN' atac ORELSE'
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
   103
    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   104
       (@{thms id_def o_def split_def sum.cases} @ map_comps @ map_idents)))));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   105
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   106
fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_fun sel_funs =
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   107
  HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   108
    (the_default (K all_tac) (Option.map rtac maybe_disc_fun)) THEN' REPEAT_DETERM_N m o atac) THEN
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   109
  unfold_thms_tac ctxt (Let_def :: sel_funs) THEN HEADGOAL (rtac refl);
53706
9e28c41e3595 more primcorec tactics
blanchet
parents: 53702
diff changeset
   110
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   111
fun inst_split_eq ctxt split =
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   112
  (case prop_of split of
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   113
    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   114
    let
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   115
      val s = Name.uu;
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   116
      val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   117
      val split' = Drule.instantiate' [] [SOME (certify ctxt eq)] split;
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   118
    in
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   119
      Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split'
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   120
    end
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   121
  | _ => split);
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   122
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   123
fun distinct_in_prems_tac distincts =
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   124
  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   125
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   126
fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m fun_ctr =
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   127
  let
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   128
    val splits' =
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   129
      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits)
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   130
  in
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   131
    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   132
    mk_primcorec_prelude ctxt [] (fun_ctr RS trans) THEN
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   133
    HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   134
      SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   135
      (rtac refl ORELSE' atac ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   136
       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   137
       Splitter.split_tac (split_if :: splits) ORELSE'
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   138
       Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   139
       mk_primcorec_assumption_tac ctxt discIs ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   140
       distinct_in_prems_tac distincts ORELSE'
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   141
       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   142
  end;
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
   143
54842
a020f7d74fed renamed 'exhaust' to 'nchotomy' since it's expressed in the object logic (cf. datatype package)
blanchet
parents: 54832
diff changeset
   144
fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'});
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   145
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   146
fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms fun_ctrs
54842
a020f7d74fed renamed 'exhaust' to 'nchotomy' since it's expressed in the object logic (cf. datatype package)
blanchet
parents: 54832
diff changeset
   147
    maybe_nchotomy =
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   148
  let
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   149
    val n = length ms;
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   150
    val (ms', fun_ctrs') =
54842
a020f7d74fed renamed 'exhaust' to 'nchotomy' since it's expressed in the object logic (cf. datatype package)
blanchet
parents: 54832
diff changeset
   151
      (case maybe_nchotomy of
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   152
        NONE => (ms, fun_ctrs)
54842
a020f7d74fed renamed 'exhaust' to 'nchotomy' since it's expressed in the object logic (cf. datatype package)
blanchet
parents: 54832
diff changeset
   153
      | SOME nchotomy =>
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   154
        (ms |> split_last ||> K [n - 1] |> op @,
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   155
         fun_ctrs
54843
7f30d569da08 made tactic work with theorems with multiple assumptions
blanchet
parents: 54842
diff changeset
   156
         |> split_last
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
   157
         ||> unfold_thms ctxt [atomize_conjL]
54908
f4ae538b0ba5 gracefully handle single-equation case, where 'nchotomy' is 'True'
blanchet
parents: 54907
diff changeset
   158
         ||> (fn thm => [rulify_nchotomy n nchotomy RS thm] handle THM _ => [thm])
54843
7f30d569da08 made tactic work with theorems with multiple assumptions
blanchet
parents: 54842
diff changeset
   159
         |> op @));
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   160
  in
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   161
    EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   162
      ms' fun_ctrs') THEN
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   163
    IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   164
      HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)))
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   165
  end;
53693
71b020d161c5 new tactics for constructor view
blanchet
parents: 53692
diff changeset
   166
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   167
fun mk_primcorec_code_of_raw_code_tac ctxt distincts splits raw =
54101
94f2dc9aea7a strengthened tactic w.r.t. "let"
blanchet
parents: 54075
diff changeset
   168
  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   169
    SELECT_GOAL (unfold_thms_tac ctxt [Let_def]) THEN' REPEAT_DETERM o
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   170
    (rtac refl ORELSE' atac ORELSE'
53904
446076262e92 got rid of dependency on silly 'eq_ifI' theorem
blanchet
parents: 53903
diff changeset
   171
     resolve_tac split_connectI ORELSE'
446076262e92 got rid of dependency on silly 'eq_ifI' theorem
blanchet
parents: 53903
diff changeset
   172
     Splitter.split_tac (split_if :: splits) ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   173
     distinct_in_prems_tac distincts ORELSE'
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   174
     rtac sym THEN' atac ORELSE'
54042
ad7a2cfb8cb2 got rid of needless argument
blanchet
parents: 54024
diff changeset
   175
     etac notE THEN' atac));
53921
46fc95abef09 tactic cleanup
blanchet
parents: 53910
diff changeset
   176
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   177
end;