src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML
author blanchet
Wed, 18 Jun 2014 17:42:24 +0200
changeset 57279 88263522c31e
parent 55990 41c6b99c5fb7
child 57399 cfc19f0a6261
permissions -rw-r--r--
made tactic more robust in the face of 'primcorec' specifications containing 'case_sum'
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_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
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
    11
  val mk_primcorec_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    12
  val mk_primcorec_ctr_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
54969
0ac0b6576d21 generate 'disc_iff' for all discriminators
blanchet
parents: 54959
diff changeset
    15
  val mk_primcorec_disc_iff_tac: Proof.context -> string list -> thm -> thm list -> thm list list ->
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    16
    thm list -> tactic
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    17
  val mk_primcorec_exhaust_tac: Proof.context -> string list -> int -> thm -> tactic
54924
44373f3560c7 refactoring
blanchet
parents: 54923
diff changeset
    18
  val mk_primcorec_nchotomy_tac: Proof.context -> thm list -> tactic
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
    19
  val mk_primcorec_raw_code_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
    20
    int list -> thm list -> thm option -> tactic
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
    21
  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
54955
cf8d429dc24e reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
blanchet
parents: 54954
diff changeset
    22
    thm list -> 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
    23
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    24
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54174
diff changeset
    25
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
    26
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    27
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    28
open BNF_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    29
open BNF_Tactics
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    30
open BNF_FP_Util
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    31
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
    32
val atomize_conjL = @{thm atomize_conjL};
53905
blanchet
parents: 53904
diff changeset
    33
val falseEs = @{thms not_TrueE FalseE};
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
    34
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
    35
val split_if = @{thm split_if};
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    36
val split_if_asm = @{thm split_if_asm};
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    37
val split_connectI = @{thms allI impI conjI};
54953
a2f4cf3387fc strengthen tac w.r.t. lets with tuples
blanchet
parents: 54926
diff changeset
    38
val unfold_lets = @{thms Let_def[abs_def] split_beta}
53902
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
    39
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    40
fun exhaust_inst_as_projs ctxt frees thm =
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    41
  let
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    42
    val num_frees = length frees;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    43
    val fs = Term.add_vars (prop_of thm) [] |> filter (can dest_funT o snd);
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    44
    fun find s = find_index (curry (op =) s) frees;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    45
    fun mk_cfp (f as ((s, _), T)) =
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    46
      (certify ctxt (Var f), certify ctxt (mk_proj T num_frees (find s)));
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    47
    val cfps = map mk_cfp fs;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    48
  in
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    49
    Drule.cterm_instantiate cfps thm
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    50
  end;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    51
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    52
val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    53
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    54
fun distinct_in_prems_tac distincts =
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    55
  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    56
54924
44373f3560c7 refactoring
blanchet
parents: 54923
diff changeset
    57
fun mk_primcorec_nchotomy_tac ctxt disc_exhausts =
54925
c63223067cab strengthened tactic
blanchet
parents: 54924
diff changeset
    58
  HEADGOAL (Method.insert_tac disc_exhausts THEN' clean_blast_tac ctxt);
54924
44373f3560c7 refactoring
blanchet
parents: 54923
diff changeset
    59
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    60
fun mk_primcorec_exhaust_tac ctxt frees n nchotomy =
54844
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    61
  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
    62
    HEADGOAL (atac ORELSE'
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    63
      cut_tac nchotomy THEN'
54923
ffed2452f5f6 instantiate schematics as projections to avoid HOU trouble
blanchet
parents: 54920
diff changeset
    64
      K (exhaust_inst_as_projs_tac ctxt frees) THEN'
54905
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    65
      EVERY' (map (fn k =>
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    66
          (if k < n then etac disjE else K all_tac) THEN'
54920
8f50ad61b0a9 generalized tactic to the case of several assumptions per equation
blanchet
parents: 54919
diff changeset
    67
          REPEAT o (dtac meta_mp THEN' atac ORELSE'
8f50ad61b0a9 generalized tactic to the case of several assumptions per equation
blanchet
parents: 54919
diff changeset
    68
            etac conjE THEN' dtac meta_mp THEN' atac ORELSE'
8f50ad61b0a9 generalized tactic to the case of several assumptions per equation
blanchet
parents: 54919
diff changeset
    69
            atac))
54905
2fdec6c29eb7 don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents: 54900
diff changeset
    70
        ks))
54844
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    71
  end;
630ba4d8a206 generate exhaust from nchotomy
blanchet
parents: 54843
diff changeset
    72
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
    73
fun mk_primcorec_assumption_tac ctxt discIs =
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    74
  SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    75
      not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
53926
9fc9a59ad579 strengthened tactic
blanchet
parents: 53922
diff changeset
    76
    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
    77
    eresolve_tac falseEs ORELSE'
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
    78
    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
53929
8c5aaf557421 further strengthening of tactics
blanchet
parents: 53926
diff changeset
    79
    dresolve_tac discIs THEN' atac ORELSE'
8c5aaf557421 further strengthening of tactics
blanchet
parents: 53926
diff changeset
    80
    etac notE THEN' atac ORELSE'
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54043
diff changeset
    81
    etac disjE))));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    82
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    83
val ss_fst_snd_conv = simpset_of (ss_only @{thms fst_conv snd_conv} @{context});
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    84
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    85
fun case_atac ctxt = simp_tac (put_simpset ss_fst_snd_conv ctxt);
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    86
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    87
fun same_case_tac ctxt m =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    88
  HEADGOAL (if m = 0 then rtac TrueI
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    89
    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' case_atac ctxt) THEN' case_atac ctxt);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    90
54971
blanchet
parents: 54969
diff changeset
    91
fun different_case_tac ctxt m exclude =
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    92
  HEADGOAL (if m = 0 then
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    93
      mk_primcorec_assumption_tac ctxt []
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    94
    else
54971
blanchet
parents: 54969
diff changeset
    95
      dtac exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN'
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
    96
      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
    97
54971
blanchet
parents: 54969
diff changeset
    98
fun cases_tac ctxt k m excludesss =
blanchet
parents: 54969
diff changeset
    99
  let val n = length excludesss in
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   100
    EVERY (map (fn [] => if k = n then all_tac else same_case_tac ctxt m
54971
blanchet
parents: 54969
diff changeset
   101
        | [exclude] => different_case_tac ctxt m exclude)
blanchet
parents: 54969
diff changeset
   102
      (take k (nth excludesss (k - 1))))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   103
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   104
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   105
fun prelude_tac ctxt defs thm =
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   106
  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt unfold_lets;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   107
54971
blanchet
parents: 54969
diff changeset
   108
fun mk_primcorec_disc_tac ctxt defs disc_corec k m excludesss =
blanchet
parents: 54969
diff changeset
   109
  prelude_tac ctxt defs disc_corec THEN cases_tac ctxt k m excludesss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   110
54969
0ac0b6576d21 generate 'disc_iff' for all discriminators
blanchet
parents: 54959
diff changeset
   111
fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss
0ac0b6576d21 generate 'disc_iff' for all discriminators
blanchet
parents: 54959
diff changeset
   112
    disc_excludes =
54910
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
   113
  HEADGOAL (rtac iffI THEN'
0ec2cccbf8ad properly synchronize parallel lists
blanchet
parents: 54909
diff changeset
   114
    rtac fun_exhaust THEN'
54959
30ded82ff806 fixed 'disc_iff' tactic in the case where different equations use different variable names for the function arguments
blanchet
parents: 54955
diff changeset
   115
    K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN'
54917
a426e38a8a7e made tactic handle corner cases, where some of the 'disc' properties are missing, correctly
blanchet
parents: 54916
diff changeset
   116
    EVERY' (map (fn [] => etac FalseE
54969
0ac0b6576d21 generate 'disc_iff' for all discriminators
blanchet
parents: 54959
diff changeset
   117
        | fun_discs' as [fun_disc'] =>
0ac0b6576d21 generate 'disc_iff' for all discriminators
blanchet
parents: 54959
diff changeset
   118
          if eq_list Thm.eq_thm (fun_discs', fun_discs) then
54917
a426e38a8a7e made tactic handle corner cases, where some of the 'disc' properties are missing, correctly
blanchet
parents: 54916
diff changeset
   119
            REPEAT_DETERM o etac conjI THEN' (atac ORELSE' rtac TrueI)
a426e38a8a7e made tactic handle corner cases, where some of the 'disc' properties are missing, correctly
blanchet
parents: 54916
diff changeset
   120
          else
54977
63a5e0d8ce3b strengthened tactic to handle 'disc_iff' equations of the form '... = False'
blanchet
parents: 54971
diff changeset
   121
            rtac FalseE THEN'
63a5e0d8ce3b strengthened tactic to handle 'disc_iff' equations of the form '... = False'
blanchet
parents: 54971
diff changeset
   122
            (rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT_DETERM o atac ORELSE'
63a5e0d8ce3b strengthened tactic to handle 'disc_iff' equations of the form '... = False'
blanchet
parents: 54971
diff changeset
   123
             cut_tac fun_disc') THEN'
54918
790362e13e0d made tactic more robust
blanchet
parents: 54917
diff changeset
   124
            dresolve_tac disc_excludes THEN' etac notE THEN' atac)
54917
a426e38a8a7e made tactic handle corner cases, where some of the 'disc' properties are missing, correctly
blanchet
parents: 54916
diff changeset
   125
      fun_discss) THEN'
54977
63a5e0d8ce3b strengthened tactic to handle 'disc_iff' equations of the form '... = False'
blanchet
parents: 54971
diff changeset
   126
    (etac FalseE ORELSE'
63a5e0d8ce3b strengthened tactic to handle 'disc_iff' equations of the form '... = False'
blanchet
parents: 54971
diff changeset
   127
     resolve_tac (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
   128
54955
cf8d429dc24e reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
blanchet
parents: 54954
diff changeset
   129
fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_idents map_comps fun_sel k
54971
blanchet
parents: 54969
diff changeset
   130
    m excludesss =
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   131
  prelude_tac ctxt defs (fun_sel RS trans) THEN
54971
blanchet
parents: 54969
diff changeset
   132
  cases_tac ctxt k m excludesss THEN
55341
3d2c97392e25 adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
blanchet
parents: 55067
diff changeset
   133
  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE'
53905
blanchet
parents: 53904
diff changeset
   134
    eresolve_tac falseEs ORELSE'
53902
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
   135
    resolve_tac split_connectI ORELSE'
396999552212 use standard "split" properties instead of ad hoc "eq_...I"
blanchet
parents: 53901
diff changeset
   136
    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
   137
    Splitter.split_tac (split_if :: splits) ORELSE'
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53908
diff changeset
   138
    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
   139
    etac notE THEN' atac ORELSE'
55067
a452de24a877 tuned names
blanchet
parents: 55061
diff changeset
   140
    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def
57279
88263522c31e made tactic more robust in the face of 'primcorec' specifications containing 'case_sum'
blanchet
parents: 55990
diff changeset
   141
         sum.case sum.sel sum.distinct[THEN eq_False[THEN iffD2]]} @
88263522c31e made tactic more robust in the face of 'primcorec' specifications containing 'case_sum'
blanchet
parents: 55990
diff changeset
   142
       mapsx @ map_comps @ map_idents))) ORELSE'
55341
3d2c97392e25 adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
blanchet
parents: 55067
diff changeset
   143
    fo_rtac @{thm cong} ctxt ORELSE'
55990
41c6b99c5fb7 more antiquotations;
wenzelm
parents: 55642
diff changeset
   144
    rtac @{thm ext}));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   145
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   146
fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   147
  HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
54926
28b621fce2f9 more SML-ish (less Haskell-ish) naming convention
blanchet
parents: 54925
diff changeset
   148
    (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
55004
96dfb73bb11a repaired 'ctr' tactic w.r.t. 'split'
blanchet
parents: 54978
diff changeset
   149
  unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);
53706
9e28c41e3595 more primcorec tactics
blanchet
parents: 53702
diff changeset
   150
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   151
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
   152
  (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
   153
    @{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
   154
    let
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   155
      val s = Name.uu;
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   156
      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
   157
      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
   158
    in
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   159
      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
   160
    end
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   161
  | _ => split);
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   162
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
   163
fun raw_code_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
   164
  let
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   165
    val splits' =
55342
blanchet
parents: 55341
diff changeset
   166
      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits);
54138
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   167
  in
c7119e1cde3e avoid confusion between outer ifs and inner ifs, that caused tactic to go berserk
blanchet
parents: 54133
diff changeset
   168
    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   169
    prelude_tac 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
   170
    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
   171
      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
   172
      (rtac refl ORELSE' atac ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   173
       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   174
       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
   175
       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
   176
       mk_primcorec_assumption_tac ctxt discIs ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   177
       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
   178
       (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
   179
  end;
53903
27fd72593624 more powerful/robust tactics
blanchet
parents: 53902
diff changeset
   180
54842
a020f7d74fed renamed 'exhaust' to 'nchotomy' since it's expressed in the object logic (cf. datatype package)
blanchet
parents: 54832
diff changeset
   181
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
   182
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
   183
fun mk_primcorec_raw_code_tac ctxt distincts discIs splits split_asms ms fun_ctrs nchotomy_opt =
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   184
  let
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   185
    val n = length ms;
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   186
    val (ms', fun_ctrs') =
54926
28b621fce2f9 more SML-ish (less Haskell-ish) naming convention
blanchet
parents: 54925
diff changeset
   187
      (case nchotomy_opt of
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   188
        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
   189
      | SOME nchotomy =>
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   190
        (ms |> split_last ||> K [n - 1] |> op @,
54900
136fe16e726a generate 'disc_iff' property in 'primcorec'
blanchet
parents: 54899
diff changeset
   191
         fun_ctrs
54843
7f30d569da08 made tactic work with theorems with multiple assumptions
blanchet
parents: 54842
diff changeset
   192
         |> split_last
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54844
diff changeset
   193
         ||> unfold_thms ctxt [atomize_conjL]
54908
f4ae538b0ba5 gracefully handle single-equation case, where 'nchotomy' is 'True'
blanchet
parents: 54907
diff changeset
   194
         ||> (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
   195
         |> op @));
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   196
  in
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
   197
    EVERY (map2 (raw_code_single_tac ctxt distincts discIs splits split_asms) ms' fun_ctrs') THEN
54832
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   198
    IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   199
      HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)))
789fbbc092d2 implemented 'exhaustive' option in tactic
blanchet
parents: 54613
diff changeset
   200
  end;
53693
71b020d161c5 new tactics for constructor view
blanchet
parents: 53692
diff changeset
   201
54978
afc156c7e4f7 cope gracefully with missing ctr equations by plugging in 'False ==> ...'
blanchet
parents: 54977
diff changeset
   202
fun mk_primcorec_code_tac ctxt distincts splits raw =
54101
94f2dc9aea7a strengthened tactic w.r.t. "let"
blanchet
parents: 54075
diff changeset
   203
  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
54954
a4ef9253a0b8 strengthened tactics w.r.t. 'lets' and tuples
blanchet
parents: 54953
diff changeset
   204
    SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   205
    (rtac refl ORELSE' atac ORELSE'
53904
446076262e92 got rid of dependency on silly 'eq_ifI' theorem
blanchet
parents: 53903
diff changeset
   206
     resolve_tac split_connectI ORELSE'
446076262e92 got rid of dependency on silly 'eq_ifI' theorem
blanchet
parents: 53903
diff changeset
   207
     Splitter.split_tac (split_if :: splits) ORELSE'
54174
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   208
     distinct_in_prems_tac distincts ORELSE'
c6291ae7cd18 strengthened tactic
blanchet
parents: 54164
diff changeset
   209
     rtac sym THEN' atac ORELSE'
54042
ad7a2cfb8cb2 got rid of needless argument
blanchet
parents: 54024
diff changeset
   210
     etac notE THEN' atac));
53921
46fc95abef09 tactic cleanup
blanchet
parents: 53910
diff changeset
   211
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   212
end;