src/HOL/Tools/BNF/bnf_tactics.ML
author wenzelm
Mon, 15 Feb 2016 14:55:44 +0100
changeset 62337 d3996d5873dd
parent 60752 b48830b670a1
child 63399 d1742d1b7f0f
permissions -rw-r--r--
proper syntax;
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_tactics.ML
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
General tactics for bounded natural functors.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
signature BNF_TACTICS =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
sig
54008
b15cfc2864de refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents: 53692
diff changeset
    11
  include CTR_SUGAR_GENERAL_TACTICS
49213
975ccb0130cb some work on coiter tactic
blanchet
parents: 49125
diff changeset
    12
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    13
  val fo_rtac: Proof.context -> thm -> int -> tactic
58332
be0f5d8d511b imported patch phantoms
blanchet
parents: 55990
diff changeset
    14
  val subst_tac: Proof.context -> int list option -> thm list -> int -> tactic
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58332
diff changeset
    15
  val subst_asm_tac: Proof.context -> int list option -> thm list -> int -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    17
  val mk_rotate_eq_tac: Proof.context -> (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list ->
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    18
    ''a list -> int -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    20
  val mk_pointfree: Proof.context -> thm -> thm
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    21
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    22
  val mk_Abs_bij_thm: Proof.context -> thm -> thm -> thm
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    23
  val mk_Abs_inj_thm: thm -> thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
55756
565a20b22f09 made tactics more robust
traytel
parents: 55569
diff changeset
    25
  val mk_map_comp_id_tac: Proof.context -> thm -> tactic
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
    26
  val mk_map_cong0_tac: Proof.context -> int -> thm -> tactic
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    27
  val mk_map_cong0L_tac: Proof.context -> int -> thm -> thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
structure BNF_Tactics : BNF_TACTICS =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
54008
b15cfc2864de refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents: 53692
diff changeset
    33
open Ctr_Sugar_General_Tactics
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
55341
3d2c97392e25 adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
blanchet
parents: 55197
diff changeset
    36
(*stolen from Christian Urban's Cookbook (and adapted slightly)*)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    37
fun fo_rtac ctxt thm = Subgoal.FOCUS (fn {concl, context = ctxt, ...} =>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    39
    val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
    val insts = Thm.first_order_match (concl_pat, concl)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    42
    rtac ctxt (Drule.instantiate_normalize insts thm) 1
55341
3d2c97392e25 adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
blanchet
parents: 55197
diff changeset
    43
  end
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    44
  handle Pattern.MATCH => no_tac) ctxt;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
58332
be0f5d8d511b imported patch phantoms
blanchet
parents: 55990
diff changeset
    46
(*unlike "unfold_thms_tac", it succeed when the RHS contains schematic variables not in the LHS*)
be0f5d8d511b imported patch phantoms
blanchet
parents: 55990
diff changeset
    47
fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt o the_default [0];
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58332
diff changeset
    48
fun subst_asm_tac ctxt = EqSubst.eqsubst_asm_tac ctxt o the_default [0];
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58332
diff changeset
    49
58332
be0f5d8d511b imported patch phantoms
blanchet
parents: 55990
diff changeset
    50
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    51
(*transforms f (g x) = h (k x) into f o g = h o k using first order matches for f, g, h, and k*)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    52
fun mk_pointfree ctxt thm = thm
58370
ffc8669e46cf made 'mk_pointfree' work again in local theories
blanchet
parents: 58352
diff changeset
    53
  |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58370
diff changeset
    54
  |> apply2 (dest_comb #> apsnd (dest_comb #> fst) #> HOLogic.mk_comp)
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    55
  |> mk_Trueprop_eq
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    56
  |> (fn goal => Goal.prove_sorry ctxt [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    57
    (K (rtac ctxt @{thm ext} 1 THEN
58370
ffc8669e46cf made 'mk_pointfree' work again in local theories
blanchet
parents: 58352
diff changeset
    58
        unfold_thms_tac ctxt ([o_apply, unfold_thms ctxt [o_apply] (mk_sym thm)]) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    59
        rtac ctxt refl 1)))
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    60
  |> Thm.close_derivation;
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    61
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    63
(* Theorems for open typedefs with UNIV as representing set *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    65
fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    66
fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac ctxt surj THEN' etac ctxt exI) 1)
55067
a452de24a877 tuned names
blanchet
parents: 55061
diff changeset
    67
  (Abs_inj_thm RS @{thm bijI'});
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
(* General tactic generators *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
(*applies assoc rule to the lhs of an equation as long as possible*)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    74
fun mk_flatten_assoc_tac ctxt refl_tac trans assoc cong = rtac ctxt trans 1 THEN
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    75
  REPEAT_DETERM (CHANGED ((FIRST' [rtac ctxt trans THEN' rtac ctxt assoc, rtac ctxt cong THEN' refl_tac]) 1)) THEN
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
  refl_tac 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
(*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
from lhs by the given permutation of monoms*)
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    80
fun mk_rotate_eq_tac ctxt refl_tac trans assoc com cong =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
    fun gen_tac [] [] = K all_tac
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
      | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
      | gen_tac (x :: xs) (y :: ys) = if x = y
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    85
        then rtac ctxt cong THEN' refl_tac THEN' gen_tac xs ys
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    86
        else rtac ctxt trans THEN' rtac ctxt com THEN'
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    87
          K (mk_flatten_assoc_tac ctxt refl_tac trans assoc cong) THEN'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
          gen_tac (xs @ [x]) (y :: ys)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
      | gen_tac _ _ = error "mk_rotate_eq_tac: different lists";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
    gen_tac
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
55756
565a20b22f09 made tactics more robust
traytel
parents: 55569
diff changeset
    94
fun mk_map_comp_id_tac ctxt map_comp0 =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    95
  (rtac ctxt trans THEN' rtac ctxt map_comp0 THEN' K (unfold_thms_tac ctxt @{thms comp_id}) THEN' rtac ctxt refl) 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55067
diff changeset
    97
fun mk_map_cong0_tac ctxt m map_cong0 =
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    98
  EVERY' [rtac ctxt mp, rtac ctxt map_cong0,
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
    99
    CONJ_WRAP' (K (rtac ctxt ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
49284
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49228
diff changeset
   100
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
   101
fun mk_map_cong0L_tac ctxt passive map_cong0 map_id =
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
   102
  (rtac ctxt trans THEN' rtac ctxt map_cong0 THEN' EVERY' (replicate passive (rtac ctxt refl))) 1 THEN
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   103
  REPEAT_DETERM (EVERY' [rtac ctxt trans, etac ctxt bspec, assume_tac ctxt,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60728
diff changeset
   104
      rtac ctxt sym, rtac ctxt @{thm id_apply}] 1) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 59582
diff changeset
   105
  rtac ctxt map_id 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
end;