src/HOL/BNF/Tools/bnf_tactics.ML
author wenzelm
Sat, 14 Dec 2013 17:28:05 +0100
changeset 54742 7a86358a3c0b
parent 54543 2d23e9c3b66b
child 54841 af71b753c459
permissions -rw-r--r--
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.; clarified tool context in some boundary cases;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49507
diff changeset
     1
(*  Title:      HOL/BNF/Tools/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
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
  val fo_rtac: thm -> Proof.context -> int -> tactic
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49463
diff changeset
    14
  val mk_unfold_thms_then_tac: Proof.context -> thm list -> ('a -> tactic) -> 'a -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
  val mk_flatten_assoc_tac: (int -> tactic) -> thm -> thm -> thm -> tactic
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
    int -> tactic
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
49518
b377da40244b renamed LFP low-level rel property to have ctor not dtor in its name
blanchet
parents: 49516
diff changeset
    25
  val mk_ctor_or_dtor_rel_tac:
b377da40244b renamed LFP low-level rel property to have ctor not dtor in its name
blanchet
parents: 49516
diff changeset
    26
    thm -> thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
49284
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49228
diff changeset
    27
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
  val mk_map_comp_id_tac: thm -> tactic
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51717
diff changeset
    29
  val mk_map_cong0_tac: int -> thm -> {prems: 'a, context: Proof.context} -> tactic
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51717
diff changeset
    30
  val mk_map_cong0L_tac: int -> thm -> thm -> tactic
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
structure BNF_Tactics : BNF_TACTICS =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
54008
b15cfc2864de refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents: 53692
diff changeset
    36
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
    37
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
(*stolen from Christian Urban's Cookbook*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
    val concl_pat = Drule.strip_imp_concl (cprop_of thm)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
    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
    44
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
    rtac (Drule.instantiate_normalize insts thm) 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
  end);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54543
diff changeset
    48
fun mk_unfold_thms_then_tac ctxt defs tac x = unfold_thms_tac ctxt defs THEN tac x;
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49456
diff changeset
    49
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    50
(*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
    51
fun mk_pointfree ctxt thm = thm
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    52
  |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    53
  |> pairself (dest_comb #> apsnd (dest_comb #> fst) #> HOLogic.mk_comp)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    54
  |> mk_Trueprop_eq
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    55
  |> (fn goal => Goal.prove_sorry ctxt [] [] goal
53588
11a77e4aa98b more robust tactic to cover another corner case (added as example);
traytel
parents: 53560
diff changeset
    56
    (K (rtac ext 1 THEN unfold_thms_tac ctxt [o_apply, mk_sym thm] THEN rtac refl 1)))
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    57
  |> Thm.close_derivation;
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 51893
diff changeset
    58
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    60
(* 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
    61
49228
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    62
fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    63
fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
e43910ccee74 open typedefs everywhere in the package
traytel
parents: 49213
diff changeset
    64
  (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
    65
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
(* General tactic generators *)
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
(*applies assoc rule to the lhs of an equation as long as possible*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
  REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
  refl_tac 1;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
(*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
    76
from lhs by the given permutation of monoms*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
fun mk_rotate_eq_tac refl_tac trans assoc com cong =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
    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
    80
      | 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
    81
      | gen_tac (x :: xs) (y :: ys) = if x = y
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
        then rtac cong THEN' refl_tac THEN' gen_tac xs ys
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
        else rtac trans THEN' rtac com THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
          K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN'
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
          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
    86
      | 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
    87
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
    gen_tac
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
49518
b377da40244b renamed LFP low-level rel property to have ctor not dtor in its name
blanchet
parents: 49516
diff changeset
    91
fun mk_ctor_or_dtor_rel_tac srel_def IJrel_defs IJsrel_defs dtor_srel {context = ctxt, prems = _} =
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    92
  unfold_thms_tac ctxt IJrel_defs THEN
49630
9f6ca87ab405 tuned tactics
traytel
parents: 49623
diff changeset
    93
  rtac (unfold_thms ctxt (IJrel_defs @ IJsrel_defs @
54543
2d23e9c3b66b killed needless thm
blanchet
parents: 54008
diff changeset
    94
    @{thms Collect_mem_eq mem_Collect_eq fst_conv snd_conv}) dtor_srel RS trans) 1 THEN
49506
aeb0cc8889f2 renamed "rel" to "srel"
blanchet
parents: 49504
diff changeset
    95
  unfold_thms_tac ctxt (srel_def ::
53560
4b5f42cfa244 removed unused/inlinable theorems
traytel
parents: 53287
diff changeset
    96
    @{thms pair_collapse Collect_mem_eq mem_Collect_eq prod.cases fst_conv snd_conv
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49456
diff changeset
    97
      split_conv}) THEN
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49456
diff changeset
    98
  rtac refl 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   100
fun mk_map_comp_id_tac map_comp0 =
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   101
  (rtac trans THEN' rtac map_comp0 THEN' REPEAT_DETERM o stac @{thm o_id} THEN' rtac refl) 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51717
diff changeset
   103
fun mk_map_cong0_tac m map_cong0 {context = ctxt, prems = _} =
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51717
diff changeset
   104
  EVERY' [rtac mp, rtac map_cong0,
49284
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49228
diff changeset
   105
    CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49228
diff changeset
   106
53285
f09645642984 renamed BNF fact
blanchet
parents: 52913
diff changeset
   107
fun mk_map_cong0L_tac passive map_cong0 map_id =
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51717
diff changeset
   108
  (rtac trans THEN' rtac map_cong0 THEN' EVERY' (replicate passive (rtac refl))) 1 THEN
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
  REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN
53285
f09645642984 renamed BNF fact
blanchet
parents: 52913
diff changeset
   110
  rtac map_id 1;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
end;